diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 71a962d..afc7aeb 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -27,7 +27,8 @@ jobs: strategy: fail-fast: false matrix: - target: [ [ 'ci-test', 'config/tests.config', 'all' ] ] + target: [ [ 'ci-test', 'config/tests.config', 'all' ] + , [ 'sha3', 'config/tests.config', 'sha3' ] ] steps: - uses: actions/checkout@v4 - name: Compile & Cache EasyCrypt diff --git a/sha3/config/tests.config b/sha3/config/tests.config new file mode 100644 index 0000000..c04ab14 --- /dev/null +++ b/sha3/config/tests.config @@ -0,0 +1,7 @@ +[default] +bin = easycrypt +report = report.log + +[test-sha3] +okdirs = proof proof/smart_counter +args = -I proof -I proof/smart_counter -p Z3 -p Alt-Ergo@2.4 diff --git a/sha3/proof/.dir-locals.el b/sha3/proof/.dir-locals.el new file mode 100644 index 0000000..650cbbf --- /dev/null +++ b/sha3/proof/.dir-locals.el @@ -0,0 +1,4 @@ +((easycrypt-mode . + ((eval . + (cl-flet ((pre (s) (concat (locate-dominating-file buffer-file-name ".dir-locals.el") s))) + (setq easycrypt-load-path `(,(pre ".") ,(pre "smart_counter")))))))) diff --git a/sha3/proof/BlockSponge.ec b/sha3/proof/BlockSponge.ec new file mode 100644 index 0000000..2000adc --- /dev/null +++ b/sha3/proof/BlockSponge.ec @@ -0,0 +1,148 @@ +(*-------------------- Padded Block Sponge Construction ----------------*) + +require import AllCore Int Real List. +require (*--*) IRO Indifferentiability Gconcl. +require import Common SLCommon. + +(*------------------------- Indifferentiability ------------------------*) + +clone include Indifferentiability with + type p <- block * capacity, + type f_in <- block list * int, + type f_out <- block list + + rename + [module] "Indif" as "Experiment" + [module] "GReal" as "RealIndif" + [module] "GIdeal" as "IdealIndif". + +(*------------------------- Ideal Functionality ------------------------*) + +clone import IRO as BIRO with + type from <- block list, + type to <- block, + op valid <- valid_block, + op dto <- bdistr. + +(*------ Validity and Parsing/Formatting of Functionality Queries ------*) + +op format (p : block list) (n : int) = p ++ nseq (n - 1) b0. +op parse: block list -> (block list * int). + +axiom formatK bs: format (parse bs).`1 (parse bs).`2 = bs. +axiom parseK p n: 0 < n => valid_block p => parse (format p n) = (p,n). +axiom parse_nil: parse [] = ([],0). + +lemma parse_injective: injective parse. +proof. +by move=> bs1 bs2 eq_format; rewrite -formatK eq_format (@formatK bs2). +qed. + +lemma parse_valid p: valid_block p => parse p = (p,1). +proof. +move=>h;have{1}->:p=format p 1;2:smt(parseK). +by rewrite/format/=nseq0 cats0. +qed. + +(*---------------------------- Restrictions ----------------------------*) + +(** The counter for the functionnality counts the number of times the + underlying primitive is called inside the functionality. This + number is equal to the sum of the number of blocks the input + message contains and the number of additional blocks the squeezing + phase has to output. + *) + +module C = { + var c : int + proc init() = { + c <- 0; + } +}. + +module FC (F : DFUNCTIONALITY) = { + proc init () : unit = {} + + proc f (bl : block list, nb : int) = { + var z : block list <- []; + if (C.c + size bl + (max (nb - 1) 0) <= max_size) { + C.c <- C.c + size bl + (max (nb - 1) 0); + z <@ F.f(bl, nb); + } + return z; + } +}. + +module PC (P : DPRIMITIVE) = { + proc init() = {} + + proc f (a : state) = { + var z : state <- (b0, c0); + if (C.c + 1 <= max_size) { + z <@ P.f(a); + C.c <- C.c + 1; + } + return z; + } + + proc fi (a : state) = { + var z : state <- (b0, c0); + if (C.c + 1 <= max_size) { + z <@ P.fi(a); + C.c <- C.c + 1; + } + return z; + } +}. + +module DRestr (D : DISTINGUISHER) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish () : bool = { + var b : bool; + C.init(); + b <@ D(FC(F), PC(P)).distinguish(); + return b; + } +}. + + +(*----------------------------- Simulator ------------------------------*) + +module Last (F : DFUNCTIONALITY) : SLCommon.DFUNCTIONALITY = { + proc init() = {} + proc f (p : block list) : block = { + var z : block list <- []; + z <@ F.f(parse p); + return last b0 z; + } +}. + +module (Sim : SIMULATOR) (F : DFUNCTIONALITY) = Gconcl.S(Last(F)). + +(*------------------------- Sponge Construction ------------------------*) + +module (Sponge : CONSTRUCTION) (P : DPRIMITIVE) : FUNCTIONALITY = { + proc init() = {} + + proc f(xs : block list, n : int) : block list = { + var z <- []; + var (sa, sc) <- (b0, Capacity.c0); + var i <- 0; + + if (valid_block xs) { + (* absorption *) + while (xs <> []) { + (sa, sc) <@ P.f(sa +^ head b0 xs, sc); + xs <- behead xs; + } + (* squeezing *) + while (i < n) { + z <- rcons z sa; + i <- i + 1; + if (i < n) { + (sa, sc) <@ P.f(sa, sc); + } + } + } + return z; + } +}. diff --git a/sha3/proof/Common.ec b/sha3/proof/Common.ec new file mode 100644 index 0000000..e836d4b --- /dev/null +++ b/sha3/proof/Common.ec @@ -0,0 +1,733 @@ +(*------------------- Common Definitions and Lemmas --------------------*) +require import Core Int IntDiv Real List Distr. +require import Ring StdRing StdOrder StdBigop BitEncoding DProd. +require (*--*) FinType BitWord PRP Monoid. +(*---*) import IntID IntOrder Bigint Bigint.BIA IntDiv. + +pragma +implicits. + + +(* -------------------------------------------------------------------- *) +op r : { int | 2 <= r } as ge2_r. +op c : { int | 0 < c } as gt0_c. + +(* -------------------------------------------------------------------- *) + +lemma gt0_r : 0 < r. +proof. by apply/(ltr_le_trans 2)/ge2_r. qed. + +lemma ge0_r : 0 <= r. +proof. by apply/ltrW/gt0_r. qed. + +lemma ge0_c : 0 <= c. +proof. by apply/ltrW/gt0_c. qed. + +(* -------------------------------------------------------------------- *) +clone export BitWord as Capacity with + op n <- c + proof gt0_n by apply/gt0_c + + rename "word" as "capacity" + "Word" as "Capacity" + (* "dunifin" as "cdistr" *) + "zerow" as "c0". +export Capacity DCapacity. + +clone export BitWord as Block with + op n <- r + proof gt0_n by apply/gt0_r + + rename "word" as "block" + "Word" as "Block" + (* "dunifin" as "bdistr" *) + "zerow" as "b0". +export DBlock. + +op cdistr = DCapacity.dunifin. +op bdistr = DBlock.dunifin. + + +(* ------------------------- Auxiliary Lemmas ------------------------- *) + +lemma dvdz_close (n : int) : + r %| n => 0 < n < 2 * r => n = r. +proof. +move=> dvd_rn. +have [m] <- : exists m, m * r = n + by exists (n %/ r); by rewrite divzK. +move=> [gt0_m_tim_r m_tim_r_lt_2r]. +case: (m = 1)=> // /ltr_total [/ltz1 le0_m | gt1_m]. +rewrite pmulr_lgt0 1:gt0_r in gt0_m_tim_r. +have // : 0 < 0 by rewrite (@ltr_le_trans m). +rewrite ltr_pmul2r 1:gt0_r in m_tim_r_lt_2r. +rewrite -lez_add1r /= in gt1_m. +have // : 2 < 2 by rewrite (@ler_lt_trans m). +qed. + +lemma chunk_nil' ['a] r : BitChunking.chunk r [<:'a>] = []. +proof. by rewrite /chunk /= mkseq0. qed. + +lemma chunk_sing' r (xs : bool list) : + 0 < r => size xs = r => BitChunking.chunk r xs = [xs]. +proof. +move=> gt0_r sz_xs_eq_r. +by rewrite /bits2blocks /chunk sz_xs_eq_r divzz ltr0_neq0 1:gt0_r b2i1 + mkseq1 /= drop0 -sz_xs_eq_r take_size. +qed. + +lemma b0P : b0 = mkblock (nseq r false). +proof. +rewrite blockP=> i ge0_i_ltr; rewrite offunifE ge0_i_ltr /= getE ge0_i_ltr /=. +rewrite ofblockK 1:size_nseq 1:/#. +by rewrite nth_nseq. +qed. + +lemma bits2w_inj_eq (cs ds : bool list) : + size cs = r => size ds = r => mkblock cs = mkblock ds <=> cs = ds. +proof. by move=> s_cs_r s_ds_r; split=> //=; exact/mkblock_pinj. qed. + +lemma last_drop_all_but_last (y : 'a, xs : 'a list) : + xs = [] \/ drop (size xs - 1) xs = [last y xs]. +proof. +elim xs=> // z zs ih /=. +case (size zs <= 0)=> [le0_sz_zs | gt0_sz_zs]. +have sz_zs_eq0 : size zs = 0 + by rewrite (@ler_asym (size zs) 0); split=> // _; rewrite size_ge0. +by have -> : zs = [] by rewrite -size_eq0. +case (zs = [])=> // zs_non_nil. elim ih=> //. +by rewrite addzC (@last_nonempty y z). +qed. + +(*------------------------------ Primitive -----------------------------*) +clone export PRP as PRPt with + type D <- block * capacity. + +clone export StrongPRP as PRPSec. + +clone export RP as Perm with + op dD <- bdistr `*` cdistr + rename + [module type] "PRP" as "PRIMITIVE" + [module] "RP" as "Perm" + proof dD_ll. +realize dD_ll. +by apply/dprod_ll; rewrite Block.DBlock.dunifin_ll Capacity.DCapacity.dunifin_ll. +qed. + +(*---------------------- Needed Blocks Computation ---------------------*) + +lemma needed_blocks0 : (0 + r - 1) %/ r = 0. +proof. +rewrite -divz_eq0 1:gt0_r; smt(gt0_r). +qed. + +lemma needed_blocks_non_pos (n : int) : + n <= 0 => (n + r - 1) %/ r <= 0. +proof. +move=> le0_n. +rewrite (lez_trans ((r - 1) %/ r)) 1:leq_div2r 1:/# 1:ge0_r. +have -> // : (r - 1) %/ r = 0 + by rewrite -divz_eq0 1:gt0_r; smt(gt0_r). +qed. + +lemma needed_blocks_suff (n : int) : + n <= (n + r - 1) %/ r * r. +proof. +have -> : (n + r - 1) %/r * r = (n + r - 1) - (n + r - 1)%% r. ++ by rewrite {2}(@divz_eq (n + r - 1) r) #ring. +by rewrite -(@addzA n) -(@addzA n) lez_addl subz_ge0 -ltzS -(@addzA r) /= + ltz_pmod gt0_r. +qed. + +lemma needed_blocks_correct (n : int) : + 0 <= (n + r - 1) %/ r * r - n < r. +proof. +split=> [| _]. +by rewrite subz_ge0 needed_blocks_suff. +have -> : (n + r - 1) %/r * r = (n + r - 1) - (n + r - 1)%% r + by rewrite {2}(@divz_eq (n + r - 1) r) #ring. +have -> : n + r - 1 - (n + r - 1) %% r - n = r - 1 - (n + r - 1) %% r + by ring. +rewrite ltzE -(@ler_add2r (-r)) /=. +have -> : r - 1 - (n + r - 1) %% r + 1 - r = -(n + r - 1) %% r by ring. +by rewrite oppz_le0 modz_ge0 gtr_eqF 1:gt0_r. +qed. + +lemma needed_blocks_prod_r (n : int) : + (n * r + r - 1) %/ r = n. +proof. +rewrite -addzA divzMDl 1:gtr_eqF 1:gt0_r // divz_small //. +smt(gt0_r). +qed. + +lemma needed_blocks_eq_div_r (n : int) : + r %| n <=> n %/ r = (n + r - 1) %/ r. +proof. +split=> [r_dvd_n | eq_div]. +have {2}<- := divzK r n _; first trivial. +by rewrite needed_blocks_prod_r. +rewrite dvdzE. +rewrite {2}(@divz_eq n r) -!addrA @divzMDl 1:gtr_eqF 1:gt0_r // + -{1}(@addz0 (n %/ r)) in eq_div. +have eq_div_simp : (n %% r + (r - 1)) %/ r = 0 + by rewrite (@addzI (n %/ r) 0 ((n %% r + (r - 1)) %/ r)). +have [_ n_mod_r_plus_r_min1_lt_r] : 0 <= n %% r + (r - 1) < r + by rewrite divz_eq0 1:gt0_r. +have n_mod_r_plus_r_min1_lt_r_simp : n %% r <= 0 + by rewrite -(@lez_add2r (r - 1)) /= -ltzS -addzA /=. +by apply lez_anti; split=> // _; rewrite modz_ge0 1:gtr_eqF 1:gt0_r. +qed. + +lemma needed_blocks_succ_eq_div_r (n : int) : + ! r %| n <=> n %/ r + 1 = (n + r - 1) %/ r. +proof. +split=> [not_r_dvd_n | succ_eq_div]. +have {2}-> := divz_eq n r. +rewrite -!addrA divzMDl 1:gtr_eqF 1:gt0_r //; ring. +rewrite dvdzE in not_r_dvd_n. +have gt0_mod : 0 < n %% r + by rewrite ltz_def=> |>; rewrite modz_ge0 1:gtr_eqF 1:gt0_r. +have [r_le_n_mod_r_plus_r_min1 n_mod_r_plus_r_min1_lt_r] : + r <= n %% r + (r - 1) < r + r. + split=> [| _]. + by rewrite (@addrC r (-1)) addrA -{1}add0z lez_add2r -ltzS + -addrA addNz. + by rewrite (@addrC r (-1)) addrA ltz_add2r -(@ltz_add2r 1) -addrA /= + (@ltr_trans r) 1:ltz_pmod 1:gt0_r -{1}addz0 ler_lt_add 1:lezz ltr01. +have [m [-> [ge0_m lt_mr]]] : + exists (m : int), n %% r + (r - 1) = r + m /\ 0 <= m < r. + exists (n %% r + (r - 1) - r). + split; first ring. + split=> [| _]. + by rewrite -(@lez_add2r r) -addrA addNz. + by rewrite -(@ltz_add2r r) -addrA addNz. +rewrite -{1}(@mul1z r) divzMDl 1:gtr_eqF 1:gt0_r // + opprD addrA /=. +rewrite divz_small; [by rewrite ger0_norm 1:ge0_r | done]. +have not_eq_dvd : n %/ r <> (n + r - 1) %/ r by smt(). +by rewrite needed_blocks_eq_div_r. +qed. + +lemma needed_blocks_rel_div_r (n : int) : + n %/ r = (n + r - 1) %/ r \/ n %/ r + 1 = (n + r - 1) %/ r. +proof. +case: (r %| n)=> [r_dvd_n | not_r_dvd_n]. +left; by apply/needed_blocks_eq_div_r. +right; by apply/needed_blocks_succ_eq_div_r. +qed. + +(* ------------------------- Padding/Unpadding ------------------------ *) + +op num0 (n : int) = (-(n + 2)) %% r. + +op mkpad (n : int) = true :: rcons (nseq (num0 n) false) true. + +op pad (s : bool list) = s ++ mkpad (size s). + +op unpad (s : bool list) = + if !last false s then None else + let i = index true (behead (rev s)) in + if i + 1 = size s then None + else let n = size s - (i + 2) in + if i = num0 n then Some (take n s) else None. + +lemma rev_mkpad n : rev (mkpad n) = mkpad n. +proof. by rewrite /mkpad rev_cons rev_rcons rev_nseq. qed. + +lemma last_mkpad b n : last b (mkpad n) = true. +proof. by rewrite !(last_cons, last_rcons). qed. + +lemma head_mkpad b n : head b (mkpad n) = true. +proof. by []. qed. + +lemma last_pad b s : last b (pad s) = true. +proof. by rewrite last_cat last_mkpad. qed. + +lemma size_mkpad n : size (mkpad n) = num0 n + 2. +proof. +rewrite /mkpad /= size_rcons size_nseq ler_maxr. +by rewrite /num0 modz_ge0 gtr_eqF ?gt0_r. by ring. +qed. + +lemma size_pad_equiv (m : int) : + 0 <= m => m + num0 m + 2 = (m + 1) %/ r * r + r. +proof. +move=> ge0_m. +by rewrite /num0 modNz 1:/# 1:gt0_r -(@addrA _ 2) /= modzE #ring. +qed. + +lemma size_padE (s : bool list) : + size (pad s) = size s + num0 (size s) + 2. +proof. by rewrite /pad size_cat size_mkpad addrA. qed. + +lemma size_pad (s : bool list) : + size (pad s) = (size s + 1) %/ r * r + r. +proof. by rewrite size_padE size_pad_equiv 1:size_ge0. qed. + +lemma size_pad_dvd_r s : r %| size (pad s). +proof. by rewrite size_pad dvdzD 1:dvdz_mull dvdzz. qed. + +lemma dvd_r_num0 (m : int) : r %| (m + num0 m + 2). +proof. by rewrite /num0 /(%|) addrAC modzDmr subrr. qed. + +lemma num0_ge0 (m : int) : 0 <= num0 m. +proof. by rewrite /num0 modz_ge0 ?gtr_eqF ?gt0_r. qed. + +lemma num0_ltr (m : int) : num0 m < r. +proof. by rewrite /num0 ltz_pmod gt0_r. qed. + +lemma index_true_behead_mkpad n : + index true (behead (mkpad n)) = num0 n. +proof. +rewrite /mkpad -cats1 //= index_cat mem_nseq size_nseq. +by rewrite ler_maxr // /num0 modz_ge0 gtr_eqF ?gt0_r. +qed. + +lemma padE (s : bool list, n : int) : + 0 <= n < r => r %| (size s + n + 2) => + pad s = s ++ [true] ++ nseq n false ++ [true]. +proof. +move=> lt_0r dvdr; rewrite -!catA /pad /mkpad /= cats1 /num0. +by do! congr; rewrite -(dvdz_modzDr dvdr) modz_small 2:#ring /#. +qed. + +lemma padK : pcancel pad unpad. +proof. +move=> s @/unpad; rewrite last_pad /= rev_cat rev_mkpad. +pose i := index _ _. +have ^iE {1 2}->: i = (-(size s + 2)) %% r. + rewrite /i behead_cat //= index_cat {1}/mkpad /= mem_rcons /=. + by rewrite index_true_behead_mkpad. +pose b := _ = size _; case b => @/b - {b}. + rewrite modNz ?gt0_r ?ltr_spaddr ?size_ge0 //. + rewrite -(@addrA _ 2) size_pad (@addrC _ r) -!addrA => /addrI. + rewrite addrCA /= -subr_eq0 -opprD oppr_eq0 addrC -divz_eq. + by rewrite addz_neq0 ?size_ge0. +move=> sz {sz}; rewrite /num0. +have -> : size (pad s) - (i + 2) + 2 = size (pad s) - i by ring. +pose b := _ = _ %% r; case b=> @/b - {b}; last first. +have -> // : size s + 2 = size (pad s) - i + by rewrite /pad size_cat size_mkpad iE #ring. +move=> sz {sz} /=; rewrite iE -size_mkpad /pad size_cat addrK. +by rewrite take_cat /= take0 cats0. +qed. + +lemma unpadK : ocancel unpad pad. +proof. +move=> s @/unpad; case: (last false s)=> //=. +elim/last_ind: s=> //= s b ih {ih}; rewrite last_rcons => ->. +rewrite rev_rcons /= size_rcons -(inj_eq (addIr (-1))) /= ?addrK. +pose i := index _ _; case: (i = size s)=> // ne_is @/pad @/num0. +have lt_is: i < size s by rewrite ltr_neqAle ne_is -size_rev index_size. +have [ge0_i lt_siz_s_i] : 0 <= i < size s. + have le_siz_s_i : i <= size s by rewrite /i - size_rev index_size. + split=> [| _]; [rewrite index_ge0 | rewrite ltr_neqAle //]. +pose j := (size s + _ - _); case: (i = (-(j + 2)) %% r)=> iE; 2:done. (* => // iE. Loops in deploy-kms *) +apply/eq_sym; rewrite -{1}(@cat_take_drop j (rcons _ _)); congr. +have jE: j = size s - (i + 1) by rewrite /j #ring. +have [ge0_j lt_js]: 0 <= j < size s by move=> /#. +rewrite -cats1 drop_cat lt_js /= /mkpad -cats1 -cat_cons; congr=> //=. +rewrite size_take // size_cat /= ltr_spsaddr //= /num0 -iE. +have sz_js: size (drop j s) = i+1; last apply/(eq_from_nth false). ++ by rewrite size_drop //= ler_maxr ?subr_ge0 ?ltrW // /j #ring. ++ by rewrite sz_js /= addrC size_nseq ler_maxr. +rewrite sz_js => k [ge0_k lt_kSi]; rewrite nth_drop //. +move/ler_eqVlt: ge0_k => [<-|] /=. + by rewrite jE -nth_rev ?nth_index // -index_mem size_rev. +move=> lt0_k; rewrite gtr_eqF //= nth_nseq 1:/#. +have ->: j + k = (size s) - ((i-k) + 1) by rewrite /j #ring. +by rewrite -nth_rev 1:/# &(@negbRL _ true) &(before_index) /#. +qed. + +inductive unpad_spec (t : bool list) = + Unpad (s : bool list, n : int) of + (0 <= n < r) + & (r %| (size s + n + 2)) + & (t = s ++ [true] ++ nseq n false ++ [true]). + +lemma nosmt unpadP (t : bool list) : + unpad t <> None <=> unpad_spec t. +proof. +split=> [|[s n lt_nr dvd ->]]; last by rewrite -padE ?padK. +case _: (unpad t) => // s sE _. +have ->: t = pad s by rewrite -(unpadK t) sE. +apply/(Unpad s (num0 (size s))). + by rewrite num0_ge0 num0_ltr. by rewrite dvd_r_num0. +by rewrite -padE ?dvd_r_num0 // num0_ge0 num0_ltr. +qed. + +(*------------------------------ Chunking ------------------------------*) + +op chunk (bs : bool list) = BitChunking.chunk r bs. + +lemma size_chunk bs : size (chunk bs) = size bs %/ r. +proof. by apply/BitChunking.size_chunk/gt0_r. qed. + +lemma in_chunk_size bs b: mem (chunk bs) b => size b = r. +proof. by apply/BitChunking.in_chunk_size/gt0_r. qed. + +lemma chunkK bs : r %| size bs => flatten (chunk bs) = bs. +proof. by apply/BitChunking.chunkK/gt0_r. qed. + +lemma chunk_nil : chunk [] = []. +proof. by apply/(@chunk_nil' r). qed. + +lemma chunk_sing (xs : bool list) : size xs = r => chunk xs = [xs]. +proof. by apply/chunk_sing'/gt0_r. qed. + +lemma chunk_cat (xs ys : bool list) : + r %| size xs => chunk (xs ++ ys) = chunk xs ++ chunk ys. +proof. by apply/BitChunking.chunk_cat/gt0_r. qed. + +lemma chunk_padK : pcancel (chunk \o pad) (unpad \o flatten). +proof. by move=> s @/(\o); rewrite chunkK 1:size_pad_dvd_r padK. qed. + +lemma flattenK bs : + (forall b, mem bs b => size b = r) => chunk (flatten bs) = bs. +proof. by apply/BitChunking.flattenK/gt0_r. qed. + +(*--------------- Converting Between Block Lists and Bits --------------*) + +op blocks2bits (xs:block list) : bool list = flatten (map ofblock xs). + +lemma blocks2bits_nil : blocks2bits [] = []. +proof. by rewrite /blocks2bits /= flatten_nil. qed. + +lemma blocks2bits_sing (x : block) : blocks2bits [x] = ofblock x. +proof. by rewrite /blocks2bits /flatten /= cats0. qed. + +lemma blocks2bits_cat (xs ys : block list) : + blocks2bits (xs ++ ys) = blocks2bits xs ++ blocks2bits ys. +proof. by rewrite /blocks2bits map_cat flatten_cat. qed. + +lemma size_blocks2bits (xs : block list) : + size (blocks2bits xs) = r * size xs. +proof. +elim: xs=> [| x xs ih]; first by rewrite blocks2bits_nil. +rewrite -cat1s blocks2bits_cat blocks2bits_sing size_cat //. +rewrite size_cat size_block ih /= #ring. +qed. + +lemma size_blocks2bits_dvd_r (xs : block list) : r %| size(blocks2bits xs). +proof. by rewrite size_blocks2bits dvdz_mulr dvdzz. qed. + +op bits2blocks (xs : bool list) : block list = map mkblock (chunk xs). + +lemma bits2blocks_nil : bits2blocks [] = []. +proof. by rewrite /bits2blocks chunk_nil. qed. + +lemma bits2blocks_sing (xs : bool list) : + size xs = r => bits2blocks xs = [mkblock xs]. +proof. by move=> sz_xs_eq_r; rewrite /bits2blocks chunk_sing. qed. + +lemma bits2blocks_cat (xs ys : bool list) : r %| size xs => r %| size ys => + bits2blocks (xs ++ ys) = bits2blocks xs ++ bits2blocks ys. +proof. +move=> r_dvd_sz_xs r_dvd_sz_ys. +by rewrite /bits2blocks chunk_cat 2:map_cat. +qed. + +lemma blocks2bitsK : cancel blocks2bits bits2blocks. +proof. +move=> xs; rewrite /blocks2bits /bits2blocks flattenK. + by move=> b /mapP [x [_ ->]];rewrite size_block. +rewrite -map_comp -{2}(@map_id xs) /(\o) /=. +by apply eq_map=> @/idfun x /=; exact/mkblockK. +qed. + +lemma bits2blocksK (bs : bool list) : + r %| size bs => blocks2bits(bits2blocks bs) = bs. +proof. +move=> dvd_r_bs; rewrite /blocks2bits /bits2blocks -map_comp. +have map_tolistK : + forall (xss : bool list list), + (forall (xs : bool list), mem xss xs => size xs = r) => + map (ofblock \o mkblock) xss = xss. ++ elim=> [// | xs yss ih eqr_sz /=]; split. + by apply ofblockK; rewrite eqr_sz. + by apply/ih => zs mem_zss_zs; rewrite eqr_sz //=; right. +by rewrite map_tolistK; [apply in_chunk_size | exact chunkK]. +qed. + +(*-------------- Padding to Blocks / Unpadding from Blocks -------------*) + +op pad2blocks : bool list -> block list = bits2blocks \o pad. +op unpad_blocks : block list -> bool list option = unpad \o blocks2bits. + +lemma pad2blocksK : pcancel pad2blocks unpad_blocks. +proof. +move=> xs @/pad2blocks @/unpad_blocks @/(\o). +by rewrite bits2blocksK 1:size_pad_dvd_r padK. +qed. + +lemma unpadBlocksK : ocancel unpad_blocks pad2blocks. +proof. +move=> xs; rewrite /pad2blocks /unpad_blocks /(\o). +pose bs := blocks2bits xs. +case (unpad bs = None) => [-> // | unpad_bs_neq_None]. +have unpad_bs : unpad bs = Some(oget(unpad bs)) + by move: unpad_bs_neq_None; case (unpad bs). +rewrite unpad_bs /=. +have -> : pad(oget(unpad bs)) = bs + by rewrite - {2} (unpadK bs) unpad_bs. +by rewrite /bs blocks2bitsK. +qed. + +lemma pad2blocks_inj : injective pad2blocks. +proof. +apply /(pcan_inj pad2blocks unpad_blocks) /pad2blocksK. +qed. + +lemma size_pad2blocks s : + size (pad2blocks s) = (size s + 1) %/ r + 1. +proof. +rewrite /pad2blocks /bits2blocks /(\o) size_map size_chunk size_pad. +have -> : (size s + 1) %/ r * r + r = ((size s + 1) %/r + 1) * r + by rewrite mulzDl mul1r. +by rewrite mulzK 1:gtr_eqF 1:gt0_r. +qed. + +(*-------------------------- Extending/Stripping -----------------------*) + +op extend (xs : block list) (n : int) = + xs ++ nseq n b0. + +op strip (xs : block list) = + let i = find (fun x => x <> b0) (rev xs) in + (take (size xs - i) xs, i). + +lemma strip_ge0 (xs : block list) : + 0 <= (strip xs).`2. +proof. by rewrite /strip /= find_ge0. qed. + +lemma extendK (xs : block list) (n : int) : + last b0 xs <> b0 => 0 <= n => strip(extend xs n) = (xs, n). +proof. +move=> xs_ends_not_b0 ge0_n; rewrite /strip /extend /=. +rewrite rev_cat rev_nseq size_cat size_nseq ler_maxr // -addzA. +have head_rev_xs_neq_b0 : head b0 (rev xs) <> b0 by rewrite - last_rev revK //. +have -> : rev xs = head b0 (rev xs) :: behead (rev xs). + by rewrite head_behead //; case: (rev xs) head_rev_xs_neq_b0. +pose p := fun (x : block) => x <> b0. +have has_p_full : has p (nseq n b0 ++ head b0 (rev xs) :: behead (rev xs)) + by apply has_cat; right; simplify; left. +have not_has_p_nseq : ! has p (nseq n b0) by rewrite has_nseq. +have -> : find p (nseq n b0 ++ head b0 (rev xs) :: behead (rev xs)) = n. + rewrite find_cat not_has_p_nseq /= size_nseq ler_maxr //. + have -> // : p (head b0 (rev xs)) by trivial. +by rewrite (@addzC n) addNz /= take_size_cat. +qed. + +lemma stripK (xs : block list) : + extend (strip xs).`1 (strip xs).`2 = xs. +proof. +rewrite /extend /strip eq_sym /=; pose i := find _ _. +rewrite -{1}(@cat_take_drop (size xs - i) xs); congr. +have [ge0_i le_ixs]: 0 <= i <= size xs. + by rewrite find_ge0 -size_rev find_size. +have sz_drop: size (drop (size xs - i) xs) = i. + rewrite size_drop ?subr_ge0 // opprD opprK. + by rewrite addrA /= ler_maxr. +apply/(eq_from_nth b0) => [|j]; rewrite ?size_nseq ?ler_maxr //. +rewrite sz_drop=> -[ge0_j lt_ji]; rewrite nth_nseq //. +rewrite nth_drop ?subr_ge0 // -{1}revK nth_rev ?size_rev. + rewrite addr_ge0 ?subr_ge0 //= -ltr_subr_addr. + by rewrite ltr_add2l ltr_opp2. +have @/predC1 /= ->// := (before_find b0 (predC1 b0)). +pose s := (_ - _)%Int; rewrite -/i (_ : s = i - (j+1)) /s 1:#ring. +by rewrite subr_ge0 -ltzE lt_ji /= ltr_snaddr // oppr_lt0 ltzS. +qed. + +(*------------------------------ Validity ------------------------------*) + +(* in Sponge *) + +op valid_toplevel (_ : bool list) = true. + +(* in BlockSponge *) + +op valid_block (xs : block list) = unpad_blocks xs <> None. + +lemma valid_pad2blocks (bs : bool list) : + valid_block(pad2blocks bs). +proof. +by rewrite /valid_block pad2blocksK. +qed. + +inductive valid_block_spec (xs : block list) = + ValidBlock (s : bool list, n : int) of + (0 <= n < r) + & (blocks2bits xs = s ++ [true] ++ nseq n false ++ [true]). + +lemma nosmt valid_blockP (xs : block list) : + valid_block xs <=> valid_block_spec xs. +proof. +split=> [vb | [s n] [rng_n b2b] b2b_xs_eq]. +have [up _] := (unpadP (blocks2bits xs)). +rewrite vb /= in up; case: up=> [s n rng_n _ b2b]. +by apply (@ValidBlock xs s n). +rewrite /valid_block unpadP (@Unpad (blocks2bits xs) s n) //. +have <- : size (blocks2bits xs) = size s + n + 2 + by rewrite b2b_xs_eq 3!size_cat /= size_nseq ler_maxr /#ring. +by apply size_blocks2bits_dvd_r. +qed. + +lemma valid_block_ends_not_b0 (xs : block list) : + valid_block xs => last b0 xs <> b0. +proof. +move=> vb_xs; have bp := valid_blockP xs. +rewrite vb_xs /= in bp. +move: bp=> [s n] _ b2b_xs_eq. +case: (last b0 xs <> b0)=> [// | last_xs_eq_b0]. +rewrite negbK in last_xs_eq_b0. +have xs_non_nil : xs <> []. + by case: xs b2b_xs_eq last_xs_eq_b0 vb_xs. +elim (last_drop_all_but_last b0 xs)=> // drop_xs. +have xs_take_drop : xs = take (size xs - 1) xs ++ drop (size xs - 1) xs + by rewrite cat_take_drop. +rewrite drop_xs last_xs_eq_b0 b0P in xs_take_drop. +have last_b2b_xs_true : last true (blocks2bits xs) = true + by rewrite b2b_xs_eq cats1 last_rcons. +have last_b2b_xs_false : last true (blocks2bits xs) = false + by rewrite xs_take_drop blocks2bits_cat blocks2bits_sing ofblockK + 1:size_nseq 1:ler_maxr 1:ge0_r // last_cat + last_nseq 1:gt0_r. +by rewrite last_b2b_xs_true in last_b2b_xs_false. +qed. + +inductive valid_block_struct_spec (xs : block list) = + ValidBlockStruct1 (ys : block list, x : block, s : bool list, n : int) of + (xs = ys ++ [x]) + & (0 <= n) + & (ofblock x = s ++ [true] ++ nseq n false ++ [true]) +| ValidBlockStruct2 (ys : block list, y z : block) of + (xs = ys ++ [y; z]) + & (last false (ofblock y)) + & (ofblock z = nseq (r - 1) false ++ [true]). + +lemma nosmt valid_block_structP (xs : block list) : + valid_block xs <=> valid_block_struct_spec xs. +proof. +rewrite valid_blockP. +split=> [[s n] [ge0_n lt_nr] b2b_xs_eq | + [ys x s n xs_eq ge0_n w2b_x_eq | + ys y z xs_eq lst w2b_z_eq]]. +have sz_s_divz_eq : size s = size s %/ r * r + size s %% r + by apply divz_eq. +pose tke := take (size s %/ r * r) s; pose drp := drop (size s %/ r * r) s. +have sz_tke : size tke = size s %/ r * r. + rewrite size_take 1:mulr_ge0 1:divz_ge0 1:gt0_r 1:size_ge0 + 1:ge0_r. + case (size s %/ r * r < size s)=> // not_lt_sz_s. + rewrite -lezNgt in not_lt_sz_s; apply ler_asym; split=> // _. + by rewrite lez_floor gtr_eqF 1:gt0_r //. +have sz_drp : size drp = size s %% r. + rewrite size_drop 1:mulr_ge0 1:divz_ge0 1:gt0_r 1:size_ge0 + 1:ge0_r. + case (size s %/ r * r < size s)=> // not_lt_sz_s. + rewrite ler_maxr /#. + have eq : size s %/ r * r = size s. + rewrite -lezNgt in not_lt_sz_s; apply ler_asym; split=> //. + by rewrite lez_floor gtr_eqF 1:gt0_r //. + rewrite ler_maxl /#. +have sz_s_pad_dvd_r : r %| (size s + n + 2). + have <- : size (s ++ [true] ++ nseq n false ++ [true]) = size s + n + 2 + by rewrite !size_cat /= size_nseq ler_maxr 1:ge0_n #ring. + rewrite -b2b_xs_eq size_blocks2bits_dvd_r. +have sz_tke_dvd_r : r %| size tke by rewrite sz_tke dvdz_mull dvdzz. +have sz_drp_plus_n_plus_2_dvd_r : r %| (size drp + n + 2). + rewrite sz_drp dvdzE + -(@dvdz_modzDl (size s %/ r * r) (size s %% r + n + 2) r) + 1:dvdz_mull 1:dvdzz. + have -> : size s %/ r * r + (size s %% r + n + 2) = size s + n + 2. + rewrite {3}sz_s_divz_eq #ring. by rewrite -dvdzE. +have xs_eq : xs = bits2blocks(s ++ [true] ++ nseq n false ++ [true]) + by rewrite -blocks2bitsK b2b_xs_eq. +rewrite -(@cat_take_drop (size s %/ r * r) s) -!catA -/tke -/drp + bits2blocks_cat in xs_eq. ++ rewrite sz_tke_dvd_r. rewrite !size_cat /= size_nseq ler_maxr 1:ge0_n. ++ have -> : size drp + (1 + (n + 1)) = size drp + n + 2 by ring. ++ rewrite sz_drp_plus_n_plus_2_dvd_r. +case: (n = r - 1)=> [n_eq_r_min1 | n_neq_r_min1]. +have sz_drp_plus1_dvd_r : r %| (size drp + 1). + rewrite dvdzE -(@addz0 (size drp + 1)) -{1}(@modzz r). + have {1}-> : r = n + 1 by smt(). + rewrite modzDmr. + have -> : size drp + 1 + (n + 1) = size drp + n + 2 by ring. + by rewrite -dvdzE. +have sz_drp_plus1_eq_r : size drp + 1 = r. + rewrite (@dvdz_close (size drp + 1)) //. + split=> [| _]; first rewrite ltr_paddl 1:size_ge0 ltr01. + have -> : 2 * r = r + r by ring. + rewrite ltr_add // 1:sz_drp 1:ltz_pmod 1:gt0_r ltzE ge2_r. +apply (@ValidBlockStruct2 xs (bits2blocks tke) + (mkblock (drp ++ [true])) (mkblock (nseq n false ++ [true]))). +rewrite xs_eq (@catA drp [true]) bits2blocks_cat 1:size_cat // + 1:size_cat 1:size_nseq 1:ler_maxr 1:ge0_n /= 1:/# + (@bits2blocks_sing (drp ++ [true])) 1:size_cat // + (@bits2blocks_sing (nseq n false ++ [true])) + 1:size_cat 1:size_nseq /= 1:ler_maxr 1:ge0_n /#. +rewrite ofblockK 1:size_cat //= cats1 last_rcons. +rewrite n_eq_r_min1 ofblockK 1:size_cat //= size_nseq ler_maxr /#. +have lt_n_r_min1 : n < r - 1 by smt(). +move: xs_eq. +have sz_drp_plus_n_plus_2_eq_r : size drp + n + 2 = r. + rewrite (@dvdz_close (size drp + n + 2)) // sz_drp. + have n_plus2_rng : 2 <= n + 2 <= r by smt(). + rewrite -addrA; split=> [| _]. + rewrite ltr_paddl 1:modz_ge0 1:gtr_eqF 1:gt0_r // /#. + have ->: 2 * r = r + r by ring. + have -> : r + r = (r - 1) + (r + 1) by ring. + rewrite ler_lt_add 1:-ltzS 1:-addrA /= 1:ltz_pmod 1:gt0_r. + by rewrite -(@ltr_add2r (-2)) -2!addrA. +move=> xs_eq. +rewrite (@bits2blocks_sing + (drp ++ ([true] ++ (nseq n false ++ [true])))) + in xs_eq. ++ rewrite !size_cat /= size_nseq ler_maxr 1:ge0_n 1:sz_drp. ++ have -> : size s %% r + (1 + (n + 1)) = size s %%r + n + 2 by ring. ++ by rewrite -sz_drp. +apply (@ValidBlockStruct1 xs (bits2blocks tke) + (mkblock (drp ++ ([true] ++ (nseq n false ++ [true])))) + drp n)=> //. +by rewrite ofblockK 1:!size_cat /= 1:size_nseq 1:ler_maxr 1:ge0_n + 1:-sz_drp_plus_n_plus_2_eq_r 1:#ring -!catA cat1s. +have sz_w2b_x_eq_r : size (ofblock x) = r by apply size_block. +rewrite w2b_x_eq !size_cat /= size_nseq ler_maxr // in sz_w2b_x_eq_r. +have lt_nr : n < r by smt(size_ge0). +apply (@ValidBlock xs (blocks2bits ys ++ s) n)=> //. +by rewrite xs_eq blocks2bits_cat blocks2bits_sing w2b_x_eq -!catA. +move: xs_eq. have -> : [y; z] = [y] ++ [z] by trivial. move=> xs_eq. +have w2bits_y_eq : ofblock y = take (r - 1) (ofblock y) ++ [true]. + rewrite -{1}(@cat_take_drop (r - 1) (ofblock y)); congr. + elim (last_drop_all_but_last false (ofblock y))=> + [w2b_y_nil | drop_w2b_y_last]. + have not_lst_w2b_y : ! last false (ofblock y) by rewrite w2b_y_nil. + by rewrite w2b_y_nil. + rewrite lst in drop_w2b_y_last. + by rewrite -drop_w2b_y_last size_block. +apply (@ValidBlock xs (blocks2bits ys ++ (take (r - 1) (ofblock y))) + (r - 1)). +smt(ge2_r). +rewrite xs_eq 2!blocks2bits_cat 2!blocks2bits_sing -!catA; congr. +by rewrite {1}w2bits_y_eq -catA w2b_z_eq. +qed. + +(* in AbsorbSponge *) + +op valid_absorb (xs : block list) = valid_block((strip xs).`1). + +inductive valid_absorb_spec (xs : block list) = + ValidAbsorb (ys : block list, n : int) of + (valid_block ys) + & (0 <= n) + & (xs = ys ++ nseq n b0). + +lemma nosmt valid_absorbP (xs : block list) : + valid_absorb xs <=> valid_absorb_spec xs. +proof. +rewrite /valid_absorb; split=> [strp_xs_valid | [ys n] ge0_n vb_ys ->]. +by rewrite (@ValidAbsorb xs (strip xs).`1 (strip xs).`2) + 2:(@strip_ge0 xs) // -/(extend (strip xs).`1 (strip xs).`2) (@stripK xs). +by rewrite -/(extend ys n) extendK 1:valid_block_ends_not_b0. +qed. diff --git a/sha3/proof/IRO.eca b/sha3/proof/IRO.eca new file mode 100644 index 0000000..7c3af6a --- /dev/null +++ b/sha3/proof/IRO.eca @@ -0,0 +1,130 @@ +(* Infinite random oracle, mapping values of type [from] to infinite + sequences of values of type [to], each sampled uniformly and + independently. We obviously make it lazy. Inputs not satisfying + a validity predicate are mapped to the empty list *) + +require import Core Int Bool List FSet SmtMap. + +type to, from. + +op valid : from -> bool. +op dto : to distr. + +module type IRO = { + proc init() : unit + + (* f x, returning the first n elements of the result *) + proc f(x : from, n : int) : to list +}. + +pred prefix_closed (m : (from * int,to) fmap) = + forall x n, + (x,n) \in m => + (forall i, 0 <= i < n => + (x,i) \in m). + +pred prefix_closed' (m : (from * int,to) fmap) = + forall x n i, + (x,n) \in m => + 0 <= i < n => + (x,i) \in m. + +lemma prefix_closed_equiv m: prefix_closed m <=> prefix_closed' m. +proof. smt(). qed. + +(* official version: *) + +module IRO : IRO = { + var mp : (from * int, to) fmap + + proc init() = { + mp <- empty; + } + + proc fill_in(x, n) = { + var r; + + if ((x,n) \notin mp) { + r <$ dto; + mp.[(x,n)] <- r; + } + return oget mp.[(x,n)]; + } + + proc f(x, n) = { + var b, bs; + var i <- 0; + + bs <- []; + if (valid x) { + while (i < n) { + b <@ fill_in(x, i); + bs <- rcons bs b; + i <- i + 1; + } + } + + return bs; + } +}. + +(* version for AbsorbToBlocks.ec attempt *) + +module IRO' : IRO = { + var mp : (from * int, to) fmap + var visible : (from * int) fset + + proc resample_invisible() = { + var work, x, r; + + work <- fdom mp `\` visible; + while (work <> fset0) { + x <- pick work; + r <$ dto; + mp.[x] <- r; + work <- work `\` fset1 x; + } + } + + proc init() = { + mp <- empty; + visible <- fset0; + } + + proc fill_in(x,n) = { + var r; + + if ((x,n) \notin mp) { + r <$ dto; + mp.[(x,n)] <- r; + } + return oget mp.[(x,n)]; + } + + proc f(x, n) = { + var b, bs; + var i <- 0; + + bs <- []; + if (valid x) { + visible <- visible `|` fset1 (x,n); + while (i < n) { + b <@ fill_in(x,i); + bs <- rcons bs b; + i <- i + 1; + } + } + + return bs; + } + + proc f_lazy(x, i) = { + var b <- witness; + + if (valid x /\ 0 <= i) { + visible <- visible `|` fset1 (x,i); + b <@ fill_in(x,i); + } + return b; + } +}. diff --git a/sha3/proof/IndifRO_is_secure.ec b/sha3/proof/IndifRO_is_secure.ec new file mode 100644 index 0000000..2b9d332 --- /dev/null +++ b/sha3/proof/IndifRO_is_secure.ec @@ -0,0 +1,217 @@ +require import AllCore Distr SmtMap. +require (****) SecureRO Indifferentiability. + + +type block. +type f_in. +type f_out. + +op sampleto : f_out distr. +axiom sampleto_ll : is_lossless sampleto. +axiom sampleto_fu : is_funiform sampleto. +axiom sampleto_full : is_full sampleto. + +op limit : int. +axiom limit_gt0 : 0 < limit. + +op bound : real. + + +op bound_counter : int. +axiom bound_counter_ge0 : 0 <= bound_counter. + +op increase_counter : int -> f_in -> int. +axiom increase_counter_spec c m : c <= increase_counter c m. + + +clone import SecureRO as SRO with + type from <- f_in, + type to <- f_out, + + op bound <- limit, + op sampleto <- sampleto, + op increase_counter <- increase_counter, + op bound_counter <- bound_counter + + proof * by smt(sampleto_fu sampleto_ll sampleto_full limit_gt0 + increase_counter_spec bound_counter_ge0). + + +clone import Indifferentiability as Indiff0 with + type p <- block, + type f_in <- f_in, + type f_out <- f_out. + +module RO : FUNCTIONALITY = { + proc init = SRO.RO.RO.init + proc f = SRO.RO.RO.get +}. + +module FInit (F : DFUNCTIONALITY) = { + proc init () = {} + proc get = F.f + proc f = F.f + proc set (a : f_in, b: f_out) = {} + proc sample (a: f_in) = {} + proc rem (a : f_in) = {} +}. + +module GetF (F : SRO.RO.RO) = { + proc init = F.init + proc f = F.get +}. + +module SInit (F : SRO.RO.RO) (S : SIMULATOR) = { + proc init() = { + S(GetF(F)).init(); + F.init(); + } + proc get = F.get + proc set = F.set + proc rem = F.rem + proc sample = F.sample +}. + +module FM (C : CONSTRUCTION) (P : PRIMITIVE) = { + proc init () = { + P.init(); + C(P).init(); + } + proc get = C(P).f + proc f = C(P).f + proc set (a : f_in, b: f_out) = {} + proc sample (a: f_in) = {} + proc rem (a : f_in) = {} +}. + + +module DColl (A : AdvCollision) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish = Collision(A,FInit(F)).main +}. + +section Collision. + + declare module A <: AdvCollision {-Bounder, -SRO.RO.RO, -SRO.RO.FRO}. + + declare axiom D_ll (F <: Oracle {-A}) : + islossless F.get => islossless A(F).guess. + + lemma coll_resistant_if_indifferentiable + (C <: CONSTRUCTION{-A, -Bounder}) + (P <: PRIMITIVE{-C, -A, -Bounder}) &m : + (exists (S <: SIMULATOR{-Bounder, -A}), + (forall (F <: FUNCTIONALITY), islossless F.f => islossless S(F).init) /\ + `|Pr[GReal(C,P,DColl(A)).main() @ &m : res] - + Pr[GIdeal(RO,S,DColl(A)).main() @ &m : res]| <= bound) => + Pr[Collision(A,FM(C,P)).main() @ &m : res] <= + bound + ((limit * (limit - 1) + 2)%r / 2%r * mu1 sampleto witness). + proof. + move=>[] S [] S_ll Hbound. + have->: Pr[Collision(A, FM(C,P)).main() @ &m : res] = + Pr[GReal(C, P, DColl(A)).main() @ &m : res]. + + byequiv=>//=; proc; inline*; wp; sim. + by swap {1} [1..2] 2; sim. + have/#:Pr[GIdeal(RO, S, DColl(A)).main() @ &m : res] <= + (limit * (limit - 1) + 2)%r / 2%r * mu1 sampleto witness. + have->:Pr[GIdeal(RO, S, DColl(A)).main() @ &m : res] = + Pr[Collision(A, SRO.RO.RO).main() @ &m : res]. + + byequiv=>//=; proc; inline DColl(A, RO, S(RO)).distinguish; wp; sim. + inline*; swap{2} 1 1; wp. + call{1}(S_ll RO _); auto. + by proc; auto; smt(sampleto_ll). + exact(RO_is_collision_resistant A &m). + qed. + +end section Collision. + + +module DPre (A : AdvPreimage) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + var h : f_out + proc distinguish () = { + var b; + b <@ Preimage(A,FInit(F)).main(h); + return b; + } +}. + +section Preimage. + + declare module A <: AdvPreimage {-Bounder, -SRO.RO.RO, -SRO.RO.FRO, -DPre}. + + declare axiom D_ll (F <: Oracle{-A}) : + islossless F.get => islossless A(F).guess. + + lemma preimage_resistant_if_indifferentiable + (C <: CONSTRUCTION{-A, -Bounder, -DPre}) + (P <: PRIMITIVE{-C, -A, -Bounder, -DPre}) &m hash : + (DPre.h{m} = hash) => + (exists (S <: SIMULATOR{-Bounder, -A, -DPre}), + (forall (F <: FUNCTIONALITY), islossless F.f => islossless S(F).init) /\ + `|Pr[GReal(C,P,DPre(A)).main() @ &m : res] - + Pr[GIdeal(RO,S,DPre(A)).main() @ &m : res]| <= bound) => + Pr[Preimage(A,FM(C,P)).main(hash) @ &m : res] <= + bound + (limit + 1)%r * mu1 sampleto hash. + proof. + move=>init_hash [] S [] S_ll Hbound. + have->: Pr[Preimage(A, FM(C,P)).main(hash) @ &m : res] = + Pr[GReal(C, P, DPre(A)).main() @ &m : res]. + + byequiv=>//=; proc; inline*; wp; sp; wp; sim. + by swap {2} [1..2] 4; sim; auto; smt(). + have/#:Pr[GIdeal(RO, S, DPre(A)).main() @ &m : res] <= + (limit + 1)%r * mu1 sampleto hash. + have->:Pr[GIdeal(RO, S, DPre(A)).main() @ &m : res] = + Pr[Preimage(A, SRO.RO.RO).main(hash) @ &m : res]. + + byequiv=>//=; proc; inline DPre(A, RO, S(RO)).distinguish; wp; sim. + inline*; swap{2} 1 1; wp; sim; auto. + call{1} (S_ll RO _); auto. + by proc; auto; smt(sampleto_ll). + exact(RO_is_preimage_resistant A &m hash). + qed. + +end section Preimage. + + +module D2Pre (A : AdvSecondPreimage) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + var m2 : f_in + proc distinguish () = { + var b; + b <@ SecondPreimage(A,FInit(F)).main(m2); + return b; + } +}. + +section SecondPreimage. + + declare module A <: AdvSecondPreimage {-Bounder, -SRO.RO.RO, -SRO.RO.FRO, -D2Pre}. + + declare axiom D_ll (F <: Oracle{-A}) : + islossless F.get => islossless A(F).guess. + + lemma second_preimage_resistant_if_indifferentiable + (C <: CONSTRUCTION{-A, -Bounder, -D2Pre}) + (P <: PRIMITIVE{-C, -A, -Bounder, -D2Pre}) &m mess : + (D2Pre.m2{m} = mess) => + (exists (S <: SIMULATOR{-Bounder, -A, -D2Pre}), + (forall (F <: FUNCTIONALITY), islossless F.f => islossless S(F).init) /\ + `|Pr[GReal(C,P,D2Pre(A)).main() @ &m : res] - + Pr[GIdeal(RO,S,D2Pre(A)).main() @ &m : res]| <= bound) => + Pr[SecondPreimage(A,FM(C,P)).main(mess) @ &m : res] <= + bound + (limit + 1)%r * mu1 sampleto witness. + proof. + move=>init_mess [] S [] S_ll Hbound. + have->: Pr[SecondPreimage(A, FM(C,P)).main(mess) @ &m : res] = + Pr[GReal(C, P, D2Pre(A)).main() @ &m : res]. + + byequiv=>//=; proc; inline*; wp; sp; wp; sim. + by swap {2} [1..2] 3; sim; auto; smt(). + have/#:Pr[GIdeal(RO, S, D2Pre(A)).main() @ &m : res] <= + (limit + 1)%r * mu1 sampleto witness. + have->:Pr[GIdeal(RO, S, D2Pre(A)).main() @ &m : res] = + Pr[SecondPreimage(A, SRO.RO.RO).main(mess) @ &m : res]. + + byequiv=>//=; proc; inline D2Pre(A, RO, S(RO)).distinguish; wp; sim. + inline*; swap{2} 1 1; wp; sim; auto. + call{1} (S_ll RO _); auto. + by proc; auto; smt(sampleto_ll). + exact(RO_is_second_preimage_resistant A &m mess). + qed. + +end section SecondPreimage. diff --git a/sha3/proof/Indifferentiability.eca b/sha3/proof/Indifferentiability.eca new file mode 100644 index 0000000..842756c --- /dev/null +++ b/sha3/proof/Indifferentiability.eca @@ -0,0 +1,68 @@ +(** A primitive: the building block we assume ideal **) +type p. + +module type PRIMITIVE = { + proc init(): unit + proc f(x : p): p + proc fi(x : p): p +}. + +module type DPRIMITIVE = { + proc f(x : p): p + proc fi(x : p): p +}. + +(** A functionality: the target construction **) +type f_in, f_out. + +module type FUNCTIONALITY = { + proc init(): unit + proc f(x : f_in): f_out +}. + +module type DFUNCTIONALITY = { + proc f(x : f_in): f_out +}. + +(** A construction takes a primitive and builds a functionality. + A simulator takes a functionality and simulates the primitive. + A distinguisher gets oracle access to a primitive and a + functionality and returns a boolean (its guess as to whether it + is playing with constructed functionality and ideal primitive or + with ideal functionality and simulated primitive). **) +module type CONSTRUCTION (P : DPRIMITIVE) = { + proc init() : unit {} + proc f(x : f_in): f_out { P.f } +}. + +module type SIMULATOR (F : DFUNCTIONALITY) = { + proc init() : unit { } + proc f(x : p) : p { F.f } + proc fi(x : p) : p { F.f } +}. + +module type DISTINGUISHER (F : DFUNCTIONALITY, P : DPRIMITIVE) = { + proc distinguish(): bool +}. + +module Indif (F : FUNCTIONALITY, P : PRIMITIVE, D : DISTINGUISHER) = { + proc main(): bool = { + var b; + + P.init(); + F.init(); + b <@ D(F,P).distinguish(); + return b; + } +}. + +(* Using the name Real can be a bad idea, since it can clash with the theory Real *) +module GReal(C : CONSTRUCTION, P : PRIMITIVE) = Indif(C(P),P). +module GIdeal(F : FUNCTIONALITY, S : SIMULATOR) = Indif(F,S(F)). + +(* (C <: CONSTRUCTION) applied to (P <: PRIMITIVE) is indifferentiable + from (F <: FUNCTIONALITY) if there exists (S <: SIMULATOR) such + that, for all (D <: DISTINGUISHER), + | Pr[Real(P,C,D): res] - Pr[Ideal(F,S,D): res] | is small. + We avoid the existential by providing a concrete construction for S + and the `small` by providing a concrete bound. *) diff --git a/sha3/proof/OptionIndifferentiability.eca b/sha3/proof/OptionIndifferentiability.eca new file mode 100644 index 0000000..638e9e1 --- /dev/null +++ b/sha3/proof/OptionIndifferentiability.eca @@ -0,0 +1,61 @@ +(** A primitive: the building block we assume ideal **) +type p. + +module type OPRIMITIVE = { + proc init(): unit + proc f(x : p): p option + proc fi(x : p): p option +}. + +module type ODPRIMITIVE = { + proc f(x : p): p option + proc fi(x : p): p option +}. + +(** A functionality: the target construction **) +type f_in, f_out. + +module type OFUNCTIONALITY = { + proc init(): unit + proc f(x : f_in): f_out option +}. + +module type ODFUNCTIONALITY = { + proc f(x : f_in): f_out option +}. + +(** A construction takes a primitive and builds a functionality. + A simulator takes a functionality and simulates the primitive. + A distinguisher gets oracle access to a primitive and a + functionality and returns a boolean (its guess as to whether it + is playing with constructed functionality and ideal primitive or + with ideal functionality and simulated primitive). **) +module type OCONSTRUCTION (P : ODPRIMITIVE) = { + proc init() : unit {} + proc f(x : f_in): f_out option { P.f } +}. + +module type OSIMULATOR (F : ODFUNCTIONALITY) = { + proc init() : unit { } + proc f(x : p) : p option { F.f } + proc fi(x : p) : p option { F.f } +}. + +module type ODISTINGUISHER (F : ODFUNCTIONALITY, P : ODPRIMITIVE) = { + proc distinguish(): bool +}. + +module OIndif (F : OFUNCTIONALITY, P : OPRIMITIVE, D : ODISTINGUISHER) = { + proc main(): bool = { + var b; + + P.init(); + F.init(); + b <@ D(F,P).distinguish(); + return b; + } +}. + +(* Using the name Real can be a bad idea, since it can clash with the theory Real *) +module OGReal(C : OCONSTRUCTION, P : OPRIMITIVE) = OIndif(C(P),P). +module OGIdeal(F : OFUNCTIONALITY, S : OSIMULATOR) = OIndif(F,S(F)). diff --git a/sha3/proof/SHA3Indiff.ec b/sha3/proof/SHA3Indiff.ec new file mode 100644 index 0000000..85aed88 --- /dev/null +++ b/sha3/proof/SHA3Indiff.ec @@ -0,0 +1,319 @@ +require import AllCore List IntDiv StdOrder Distr SmtMap FSet. + +require import Common Sponge. import BIRO. +require (*--*) SLCommon Gconcl_list BlockSponge. + +(* FIX: would be nicer to define limit at top-level and then clone + BlockSponge with it - so BlockSponge would then clone lower-level + theories with it + +op limit : {int | 0 < limit} as gt0_max_limit. +*) +op limit : int = SLCommon.max_size. + + + +(* The last inlined simulator *) +type state = SLCommon.state. +op parse = BlockSponge.parse. +op valid = Gconcl_list.valid. + + +module Simulator (F : DFUNCTIONALITY) = { + var m : (state, state) fmap + var mi : (state, state) fmap + var paths : (capacity, block list * block) fmap + proc init() = { + m <- empty; + mi <- empty; + paths <- empty.[c0 <- ([],b0)]; + Gconcl_list.BIRO2.IRO.init(); + } + proc f (x : state) : state = { + var p,v,z,q,k,cs,y,y1,y2; + if (x \notin m) { + if (x.`2 \in paths) { + (p,v) <- oget paths.[x.`2]; + z <- []; + (q,k) <- parse (rcons p (v +^ x.`1)); + if (valid q) { + cs <@ F.f(oget (unpad_blocks q), k * r); + z <- bits2blocks cs; + } else { + z <@ Gconcl_list.BIRO2.IRO.f(q,k); + } + y1 <- last b0 z; + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + m.[x] <- y; + mi.[y] <- x; + if (x.`2 \in paths) { + (p,v) <-oget paths.[x.`2]; + paths.[y2] <- (rcons p (v +^ x.`1),y.`1); + } + } else { + y <- oget m.[x]; + } + return y; + } + proc fi (x : state) : state = { + var y,y1,y2; + if (! x \in mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + mi.[x] <- y; + m.[y] <- x; + } else { + y <- oget mi.[x]; + } + return y; + } +}. + + + +(*---------------------------- Restrictions ----------------------------*) + +(** The counter for the functionality counts the number of times the + underlying primitive is called inside the functionality. This + number is equal to the sum of the number of blocks in the padding + of the input, plus the number of additional blocks the squeezing + phase has to output. + *) + +module Cntr = { + var c : int + + proc init() = { + c <- 0; + } +}. + +module FC (F : DFUNCTIONALITY) = { + proc init () : unit = {} + + (* ((size bs + 1) %/ r + 1) = size (pad2blocks bs): *) + + proc f (bs : bool list, n : int) : bool list = { + var z : bool list <- []; + if (Cntr.c + + ((size bs + 1) %/ r + 1) + + (max ((n + r - 1) %/ r - 1) 0) <= limit) { + Cntr.c <- + Cntr.c + + ((size bs + 1) %/ r + 1) + + (max ((n + r - 1) %/ r - 1) 0); + z <@ F.f(bs, n); + } + return z; + } +}. + +module PC (P : DPRIMITIVE) = { + proc init() = {} + + proc f (a : block * capacity) = { + var z : block * capacity <- (b0, c0); + if (Cntr.c + 1 <= limit) { + z <@ P.f(a); + Cntr.c <- Cntr.c + 1; + } + return z; + } + proc fi (a : block * capacity) = { + var z : block * capacity <- (b0, c0); + if (Cntr.c + 1 <= limit) { + z <@ P.fi(a); + Cntr.c <- Cntr.c + 1; + } + return z; + } +}. + +module DRestr (D : DISTINGUISHER) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish () : bool = { + var b : bool; + Cntr.init(); + b <@ D(FC(F),PC(P)).distinguish(); + return b; + } +}. + +section. + +declare module Dist <: + DISTINGUISHER {-Perm, -Gconcl_list.SimLast, -IRO, -Cntr, -BlockSponge.BIRO.IRO, -Simulator, -BlockSponge.C, -Gconcl.S, -SLCommon.F.RO, -SLCommon.F.FRO, -SLCommon.Redo, -SLCommon.C, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator}. + +declare axiom Dist_lossless (F <: DFUNCTIONALITY {-Dist}) (P <: DPRIMITIVE {-Dist}) : + islossless P.f => islossless P.fi => islossless F.f => + islossless Dist(F,P).distinguish. + +lemma drestr_commute1 &m : + Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, + LowerDist(DRestr(Dist))).main() @ &m : res] = + Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, + BlockSponge.DRestr(LowerDist(Dist))).main() @ &m : res]. +proof. +byequiv=> //; proc. +seq 2 2 : (={glob Dist} /\ ={Perm.m, Perm.mi} ); first sim. +inline*; wp; sp. +call (_ : ={c}(Cntr, BlockSponge.C) /\ ={Perm.m, Perm.mi}). +proc; sp; if=> //; sp; sim. +proc; sp; if=> //; sp; sim. +proc=> /=. +inline BlockSponge.FC(BlockSponge.Sponge(Perm)).f. +wp; sp. +if=> //. +progress; smt(size_pad2blocks). +seq 1 1 : + (={n} /\ nb{2} = (n{2} + r - 1) %/ r /\ bl{2} = pad2blocks bs{1} /\ + Cntr.c{1} = BlockSponge.C.c{2} /\ ={Perm.m, Perm.mi}). +auto; progress; by rewrite size_pad2blocks. +inline RaiseFun(BlockSponge.Sponge(Perm)).f. +wp; sp. +call (_ : ={Perm.m, Perm.mi}); first sim. +auto. +auto; progress; by rewrite blocks2bits_nil. +qed. + +lemma drestr_commute2 &m : + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, Gconcl_list.SimLast(Gconcl.S), + LowerDist(DRestr(Dist))).main() @ &m : res] = + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, Gconcl_list.SimLast(Gconcl.S), + BlockSponge.DRestr(LowerDist(Dist))).main() @ &m : res]. +proof. +byequiv=> //; proc. +seq 2 2 : + (={glob Dist, BlockSponge.BIRO.IRO.mp, + glob Gconcl_list.SimLast(Gconcl.S)}); first sim. +inline*; wp; sp. +call + (_ : ={BlockSponge.BIRO.IRO.mp,Gconcl_list.BIRO2.IRO.mp} /\ + ={c}(Cntr, BlockSponge.C) /\ + ={glob Gconcl_list.SimLast(Gconcl.S)}). +proc; sp; if=> //; sim. +proc; sp; if=> //; sim. +proc=> /=. +inline BlockSponge.FC(BlockSponge.BIRO.IRO).f. +sp; wp. +if=> //. +progress; smt(size_pad2blocks). +seq 1 1 : + (={n} /\ nb{2} = (n{2} + r - 1) %/ r /\ bl{2} = pad2blocks bs{1} /\ + Cntr.c{1} = BlockSponge.C.c{2} /\ + ={BlockSponge.BIRO.IRO.mp, Gconcl_list.BIRO2.IRO.mp, + Gconcl.S.paths, Gconcl.S.mi, Gconcl.S.m}). +auto; progress. +rewrite size_pad2blocks //. +inline RaiseFun(BlockSponge.BIRO.IRO).f. +wp; sp. +call (_ : ={BlockSponge.BIRO.IRO.mp}); first sim. +auto. +by auto. +qed. + +op wit_pair : block * capacity = witness. + +local equiv equiv_sim_f (F <: DFUNCTIONALITY{-Gconcl.S, -Simulator}) : + RaiseSim(Gconcl_list.SimLast(Gconcl.S),F).f + ~ + Simulator(F).f + : + ={arg, glob F, glob Gconcl_list.BIRO2.IRO} /\ ={m, mi, paths}(Gconcl.S,Simulator) + ==> + ={res, glob F, glob Gconcl_list.BIRO2.IRO} /\ ={m, mi, paths}(Gconcl.S,Simulator). +proof. +proc;inline*;if;1,3:auto=>/#. +wp;conseq(:_==> ={y1, y2, glob F, glob Gconcl_list.BIRO2.IRO} + /\ ={m, mi, paths}(Gconcl.S,Simulator));progress;sim. +if;1,3:auto=>/#;wp;sp;if;1:(auto;smt(BlockSponge.parseK BlockSponge.formatK)); + last sim;smt(BlockSponge.parseK BlockSponge.formatK). +by sp;wp;rcondt{1}1;auto;call(: true);auto;smt(BlockSponge.parseK BlockSponge.formatK). +qed. + + +local equiv equiv_sim_fi (F <: DFUNCTIONALITY{-Gconcl.S, -Simulator}) : + RaiseSim(Gconcl_list.SimLast(Gconcl.S),F).fi + ~ + Simulator(F).fi + : + ={arg, glob F, glob Gconcl_list.BIRO2.IRO} /\ ={m, mi, paths}(Gconcl.S,Simulator) + ==> + ={res, glob F, glob Gconcl_list.BIRO2.IRO} /\ ={m, mi, paths}(Gconcl.S,Simulator). +proof. by proc;inline*;if;auto=>/#. qed. + +local lemma replace_simulator &m : + Pr[IdealIndif(IRO, RaiseSim(Gconcl_list.SimLast(Gconcl.S)), + DRestr(Dist)).main() @ &m : res] = + Pr[IdealIndif(IRO, Simulator, DRestr(Dist)).main() @ &m : res]. +proof. +byequiv=>//=;proc;inline*;sp;wp. +call(: ={glob IRO, glob DRestr, glob Gconcl_list.BIRO2.IRO} + /\ ={m, mi, paths}(Gconcl.S,Simulator));auto. ++ by proc;sp;if;auto;call(equiv_sim_f IRO);auto. ++ by proc;sp;if;auto;call(equiv_sim_fi IRO);auto. +by proc;sim. +qed. + + + +lemma security &m : + `|Pr[RealIndif(Sponge, Perm, DRestr(Dist)).main() @ &m : res] - + Pr[IdealIndif(IRO, Simulator, DRestr(Dist)).main() @ &m : res]| <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + (4 * limit ^ 2)%r / (2 ^ c)%r. +proof. +rewrite -(replace_simulator &m). +rewrite exprSr 1:addz_ge0 1:ge0_r 1:ge0_c mulrC exprD_nneg 1:ge0_r 1:ge0_c. +have -> : + (limit ^ 2 - limit)%r / (2 * (2 ^ r * 2 ^ c))%r = + ((limit ^ 2 - limit)%r / 2%r) * (1%r / (2 ^ r)%r) * (1%r / (2 ^ c)%r). + by rewrite (fromintM 2); smt(). +rewrite/=. +have -> : + (4 * limit ^ 2)%r / (2 ^ c)%r = + limit%r * ((2 * limit)%r / (2 ^ c)%r) + limit%r * ((2 * limit)%r / (2 ^ c)%r). + have -> : 4 = 2 * 2 by trivial. + have {3}-> : 2 = 1 + 1 by trivial. + rewrite exprS // expr1 /#. +rewrite -/SLCommon.dstate /limit. +have->:=conclusion (Gconcl_list.SimLast(Gconcl.S)) (DRestr(Dist)) &m. +have//=:=(Gconcl_list.Real_Ideal (LowerDist(Dist)) _ &m). ++ move=>F P hp hpi hf'//=. + have hf:islossless RaiseFun(F).f. + - proc;call hf';auto. + exact(Dist_lossless (RaiseFun(F)) P hp hpi hf). +rewrite(drestr_commute1 &m) (drestr_commute2 &m). +have->:=Gconcl_list.Simplify_simulator (LowerDist(Dist)) _ &m. ++ move=>F P hp hpi hf'//=. + have hf:islossless RaiseFun(F).f. + - proc;call hf';auto. + exact(Dist_lossless (RaiseFun(F)) P hp hpi hf). +smt(). +qed. + + + + +end section. + +lemma SHA3Indiff + (Dist <: DISTINGUISHER{-Perm, -IRO, -BlockSponge.BIRO.IRO, -Cntr, -Simulator, -Gconcl_list.SimLast(Gconcl.S), -BlockSponge.C, -Gconcl.S, -SLCommon.F.RO, -SLCommon.F.FRO, -SLCommon.Redo, -SLCommon.C, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator}) + &m : + (forall (F <: DFUNCTIONALITY {-Dist}) (P <: DPRIMITIVE {-Dist}), + islossless P.f => + islossless P.fi => + islossless F.f => + islossless Dist(F,P).distinguish) => + `|Pr[RealIndif(Sponge, Perm, DRestr(Dist)).main() @ &m : res] - + Pr[IdealIndif(IRO, Simulator, DRestr(Dist)).main() @ &m : res]| <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + (4 * limit ^ 2)%r / (2 ^ c)%r. +proof. move=>h;apply (security Dist h &m). qed. + diff --git a/sha3/proof/SHA3OSecurity.ec b/sha3/proof/SHA3OSecurity.ec new file mode 100644 index 0000000..95fa0bc --- /dev/null +++ b/sha3/proof/SHA3OSecurity.ec @@ -0,0 +1,2634 @@ +(* Top-level Proof of SHA-3 Security *) + +require import AllCore Distr DList DBool List IntDiv Dexcepted DProd SmtMap FSet. +require import Common SLCommon Sponge SHA3_OIndiff. +require (****) SecureORO SecureHash. +(*****) import OIndif. + +require import PROM. + + +(* module SHA3 (P : DPRIMITIVE) = { *) +(* proc init() : unit = {} *) +(* proc f (bl : bool list, n : int) : bool list = { *) +(* var r : bool list; *) +(* r <@ Sponge(P).f(bl ++ [false; true], n); *) +(* return r; *) +(* } *) +(* }. *) + +op size_out : int. +axiom size_out_gt0 : 0 < size_out. + +op sigma : int = SHA3Indiff.limit. +axiom sigma_ge0 : 0 <= sigma. + +op limit : int = sigma. + +type f_out. + +op dout : f_out distr. +axiom dout_ll : is_lossless dout. +axiom dout_fu : is_funiform dout. +axiom dout_full : is_full dout. + + +op to_list : f_out -> bool list. +op of_list : bool list -> f_out option. +axiom spec_dout (l : f_out) : size (to_list l) = size_out. +axiom spec2_dout (l : bool list) : size l = size_out => of_list l <> None. +axiom to_list_inj : injective to_list. +axiom to_listK e l : to_list e = l <=> of_list l = Some e. + +axiom dout_equal_dlist : dmap dout to_list = dlist dbool size_out. + +lemma doutE1 x : mu1 dout x = inv (2%r ^ size_out). +proof. +have->:inv (2%r ^ size_out) = mu1 (dlist dbool size_out) (to_list x). ++ rewrite dlist1E. + - smt(size_out_gt0). + rewrite spec_dout/=. + pose p:= StdBigop.Bigreal.BRM.big _ _ _. + have->: p = StdBigop.Bigreal.BRM.big predT (fun _ => inv 2%r) (to_list x). + - rewrite /p =>{p}. + apply StdBigop.Bigreal.BRM.eq_bigr. + by move=> i; rewrite//= dbool1E. + rewrite StdBigop.Bigreal.BRM.big_const count_predT spec_dout=> {p}. + have:=size_out_gt0; move/ltzW. + move:size_out;apply intind=> //=. + - by rewrite RField.expr0 iter0 //= fromint1. + move=> i hi0 rec. + by rewrite RField.exprS//iterS// -rec; smt(). +rewrite -dout_equal_dlist dmap1E. +apply mu_eq. +by move=> l; rewrite /pred1/(\o); smt(to_listK). +qed. + + +(* module CSetSize (F : OCONSTRUCTION) (P : ODPRIMITIVE) = { *) +(* proc init = F(P).init *) +(* proc f (x : bool list) = { *) +(* var r, l; *) +(* r <@ F(P).f(x,size_out); *) +(* l <- (r <> None) ? of_list (oget r) : None; *) +(* return l; *) +(* } *) +(* }. *) + +(* module FSetSize (F : OFUNCTIONALITY) = { *) +(* proc init = F.init *) +(* proc f (x : bool list) = { *) +(* var r, l; *) +(* r <@ F.f(x,size_out); *) +(* l <- (r <> None) ? of_list (oget r) : None; *) +(* return l; *) +(* } *) +(* }. *) + +clone import SecureORO as SORO with + type from <- bool list, + type to <- f_out, + + op sampleto <- dout, + op bound <- sigma, + op increase_counter <- fun c m => c + ((size m + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 + + proof *. +realize bound_ge0 by exact(sigma_ge0). +realize sampleto_ll by exact(dout_ll). +realize sampleto_fu by exact(dout_fu). +realize sampleto_full by exact(dout_full). +realize increase_counter_spec by smt(List.size_ge0 divz_ge0 gt0_r). + + +clone import SecureHash as SH with + type from <- bool list, + type to <- f_out, + type block <- state, + op sampleto <- dout, + op bound <- sigma, + op increase_counter <- fun c m => c + ((size m + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 +proof *. +realize sampleto_ll by exact(dout_ll). +realize sampleto_fu by exact(dout_fu). +realize sampleto_full by exact(dout_full). +realize bound_ge0 by exact(sigma_ge0). +realize increase_counter_spec by smt(List.size_ge0 divz_ge0 gt0_r). + + +(* module FGetSize (F : ODFUNCTIONALITY) = { *) +(* proc f (x : bool list, i : int) = { *) +(* var r; *) +(* r <@ F.f(x); *) +(* return to_list r; *) +(* } *) +(* }. *) + +(* module SimSetSize (S : SIMULATOR) (F : Indiff0.DFUNCTIONALITY) = S(FGetSize(F)). *) + +(* module DFSetSize (F : DFUNCTIONALITY) = { *) +(* proc f (x : bool list) = { *) +(* var r; *) +(* r <@ F.f(x,size_out); *) +(* return oget (of_list r); *) +(* } *) +(* }. *) + +(* module (DSetSize (D : Indiff0.DISTINGUISHER) : DISTINGUISHER) *) +(* (F : DFUNCTIONALITY) (P : DPRIMITIVE) = D(DFSetSize(F),P). *) + + +module FSetSize (F : OFUNCTIONALITY) : OIndif.OFUNCTIONALITY = { + proc init = F.init + proc f (x : bool list) = { + var y, r; + y <@ F.f(x,size_out); + r <- (y <> None) ? of_list (oget y) : None; + return r; + } + proc get = f +}. + +module DFSetSize (F : ODFUNCTIONALITY) : OIndif.OFUNCTIONALITY = { + proc init () = {} + proc f (x : bool list) = { + var y, r; + y <@ F.f(x,size_out); + r <- (y <> None) ? of_list (oget y) : None; + return r; + } +}. + +module FIgnoreSize (F : OIndif.ODFUNCTIONALITY) : OFUNCTIONALITY = { + proc init () = {} + proc f (x : bool list, i : int) = { + var y, r; + y <@ F.f(x); + return omap to_list r; + } +}. + +module (OSponge : OIndif.OCONSTRUCTION) (P : OIndif.ODPRIMITIVE) = + FSetSize(CSome(Sponge,P)). + + +clone import Program as PBool with + type t <- bool, + op d <- dbool +proof *. + +clone import FullRO as Eager with + type in_t <- bool list * int, + type out_t <- bool, + op dout _ <- dbool, + type d_in_t <- unit, + type d_out_t <- bool. +import FullEager. + +section Preimage. + + declare module A <: SH.AdvPreimage {-Perm, -Counter, -Bounder, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -BIRO.IRO, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -SHA3Indiff.Simulator, -SHA3Indiff.Cntr, -SORO.Bounder, -SORO.RO.RO, -RO, -FRO}. + + local module FInit (F : OIndif.ODFUNCTIONALITY) : OIndif.OFUNCTIONALITY = { + proc init () = {} + proc f = F.f + }. + + local module PInit (P : ODPRIMITIVE) : OPRIMITIVE = { + proc init () = {} + proc f = P.f + proc fi = P.fi + }. + + +local module OF (F : Oracle) : OIndif.ODFUNCTIONALITY = { + proc f = F.get +}. + + +local module Log = { + var m : (bool list * int, bool) fmap +}. + +local module ExtendOutputSize (F : Oracle) : ODFUNCTIONALITY = { + proc f (x : bool list, k : int) = { + var o, l, suffix, prefix, i, r; + l <- None; + prefix <- []; + suffix <- []; + o <@ F.get(x); + prefix <- take k (to_list (oget o)); + i <- size_out; + while (i < k) { + if ((x,i) \notin Log.m) { + r <$ {0,1}; + Log.m.[(x,i)] <- r; + } + suffix <- rcons suffix (oget Log.m.[(x,i)]); + i <- i + 1; + } + l <- Some (prefix ++ suffix); + return l; + } +}. + +local module OFC2 (F : Oracle) = OFC(ExtendOutputSize(F)). + +local module ExtendOutput (F : RF) = { + proc init () = { + Log.m <- empty; + F.init(); + } + proc f = ExtendOutputSize(F).f + proc get = f +}. + + local module (Dist_of_P1Adv (A : SH.AdvPreimage) : ODISTINGUISHER) (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + proc distinguish () = { + var hash, hash', m; + Log.m <- empty; + hash <$ dout; + m <@ A(DFSetSize(F),P).guess(hash); + hash' <@ DFSetSize(F).f(m); + return hash' = Some hash; + } + }. + + +local module (SORO_P1 (A : SH.AdvPreimage) : SORO.AdvPreimage) (F : Oracle) = { + proc guess (h : f_out) : bool list = { + var mi; + Log.m <- empty; + Counter.c <- 0; + OSimulator(ExtendOutputSize(F)).init(); + mi <@ A(DFSetSize(OFC2(F)),OPC(OSimulator(ExtendOutputSize(F)))).guess(h); + return mi; + } +}. + +local module RFList = { + var m : (bool list, f_out) fmap + proc init () = { + m <- empty; + } + proc get (x : bool list) : f_out option = { + var z; + if (x \notin m) { + z <$ dlist dbool size_out; + m.[x] <- oget (of_list z); + } + return m.[x]; + } + proc sample (x: bool list) = {} +}. + +local module RFWhile = { + proc init () = { + RFList.m <- empty; + } + proc get (x : bool list) : f_out option = { + var l, i, b; + if (x \notin RFList.m) { + i <- 0; + l <- []; + while (i < size_out) { + b <$ dbool; + l <- rcons l b; + i <- i + 1; + } + RFList.m.[x] <- oget (of_list l); + } + return RFList.m.[x]; + } + proc sample (x: bool list) = {} +}. + +local equiv rw_RF_List_While : + RFList.get ~ RFWhile.get : + ={arg, glob RFList} ==> ={res, glob RFWhile}. +proof. +proc; if; 1, 3: auto; wp. +conseq(:_==> z{1} = l{2})=> />. +transitivity{1} { + z <@ Sample.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>. ++ by inline*; auto. +transitivity{1} { + z <@ LoopSnoc.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>; last first. ++ inline*; auto; sim. + by while(={l, i} /\ n{1} = size_out); auto; smt(cats1). +by call(Sample_LoopSnoc_eq); auto. +qed. + + +op inv (m1 : (bool list * int, bool) fmap) (m2 : (bool list, f_out) fmap) = + (forall l i, (l,i) \in m1 => 0 <= i < size_out) /\ + (forall l i, (l,i) \in m1 => l \in m2) /\ + (forall l, l \in m2 => forall i, 0 <= i < size_out => (l,i) \in m1) /\ + (forall l i, (l,i) \in m1 => m1.[(l,i)] = Some (nth witness (to_list (oget m2.[l])) i)). + +local equiv eq_IRO_RFWhile : + BIRO.IRO.f ~ RFWhile.get : + arg{1} = (x{2}, size_out) /\ inv BIRO.IRO.mp{1} RFList.m{2} + ==> + res{2} = of_list res{1} /\ size res{1} = size_out /\ inv BIRO.IRO.mp{1} RFList.m{2}. +proof. +proc; inline*; sp. +rcondt{1} 1; 1: by auto. +if{2}; sp; last first. ++ alias{1} 1 mp = BIRO.IRO.mp. + conseq(:_==> BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ i{1} = size_out /\ + inv mp{1} RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))=> />. + - move=> &l &r H0 H1 H2 H3 H4 bs_L mp_L H5 H6 H7 H8 H9. + rewrite take_oversize 1:spec_dout 1:H5 //. + rewrite eq_sym to_listK => ->. + by have:=H4; rewrite domE; smt(). + - smt(take_oversize spec_dout). + while{1}(BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ + 0 <= i{1} <= size_out /\ n{1} = size_out /\ + inv mp{1} RFList.m{2} /\ x{1} \in RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))(size_out - i{1}); + auto=> />. + + sp; rcondf 1; auto=> />; 1: smt(). + move=> &h H0 H1 H2 H3 H4 H5 H6 H7 H8. + rewrite size_rcons //=; do!split; 1, 2, 4: smt(size_ge0). + rewrite (take_nth witness) 1:spec_dout 1:size_ge0//=. + rewrite - H7; congr; rewrite H5=> //=. + by apply H4=> //=. + smt(size_out_gt0 size_ge0 take0). +auto=> //=. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ n{1} = size_out /\ + inv BIRO.IRO.mp{1} RFList.m{2}.[x{2} <- oget (of_list l{2})])=> />. ++ smt(get_setE spec2_dout). ++ smt(get_setE spec2_dout). +alias{1} 1 m = BIRO.IRO.mp; sp. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ move=> /> &l &r H0 H1 H2 H3 H4 mp_L bs_L H5 H6 H7 H8 H9; do!split; ..-2 : smt(domE mem_set). + move=> l j Hin. + rewrite get_setE/=. + case: (l = x{r}) => [<<-|]. + - rewrite oget_some H9; 1:smt(); congr; congr. + by rewrite eq_sym to_listK; smt(spec2_dout). + move=> Hneq. + by rewrite -(H7 _ _ Hneq) H3; smt(domE). +while(l{2} = bs{1} /\ size bs{1} = i{1} /\ 0 <= i{1} <= n{1} /\ ={i} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ sp; rcondt{1} 1; auto=> />. + - smt(). + move=> &l &r H0 H1 H2 H3 H4 H5 H6 H7 H8 H9 H10 rL _. + rewrite get_setE /= size_rcons /=; do!split; 1,2: smt(size_ge0). + - smt(mem_set). + - smt(get_setE). + - smt(mem_set). + - move=>j Hj0 Hjsize; rewrite get_setE/=nth_rcons. + case: (j = size bs{l})=>[->>//=|h]. + have/=Hjs:j < size bs{l} by smt(). + by rewrite Hjs/=H9//=. +by auto; smt(size_out_gt0). +qed. + +op eq_extend_size (m1 : (bool list * int, bool) fmap) (m2 : (bool list * int, bool) fmap) + (m3 : (bool list * int, bool) fmap) = + (* (forall x j, (x,j) \in m2 => 0 <= j < size_out) /\ *) + (* (forall x j, (x,j) \in m2 => forall k, 0 <= k < size_out => (x, k) \in m2) /\ *) + (forall x j, 0 <= j < size_out => m1.[(x,j)] = m2.[(x,j)]) /\ + (forall x j, size_out <= j => m1.[(x,j)] = m3.[(x,j)]) /\ + (forall x j, (x,j) \in m1 => 0 <= j). + +local module ExtendSample (F : OFUNCTIONALITY) = { + proc init = F.init + proc f (x : bool list, k : int) = { + var y; + if (k <= size_out) { + y <@ F.f(x,size_out); + y <- omap (take k) y; + } else { + y <@ F.f(x,k); + } + return y; + } +}. + + +local equiv eq_extend : + ExtendSample(FSome(BIRO.IRO)).f ~ ExtendOutputSize(FSetSize(FSome(BIRO.IRO))).f : + ={arg} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} ==> + ={res} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}. +proof. +proc; inline*; auto; sp. +rcondt{2} 1; 1: auto. +if{1}; sp. +- rcondt{1} 1; auto. + rcondf{2} 8; 1: auto. + - conseq(:_==> true); 1: smt(). + by while(true); auto. + auto=> /=. + conseq(:_==> ={bs, k} /\ size bs{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2})=> //=. + - smt(cats0 to_listK spec2_dout). + while(={k, bs, n, x2} /\ i{1} = i0{2} /\ n{1} = size_out /\ + 0 <= i{1} <= n{1} /\ size bs{1} = i{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> />; smt(domE get_setE size_rcons). + by auto=> />; smt(size_eq0 size_out_gt0). +rcondt{1} 1; 1: auto. +splitwhile{1} 1 : i0 < size_out; auto=> /=. +while( (i0, n0, x3){1} = (i, k, x){2} /\ bs0{1} = prefix{2} ++ suffix{2} /\ + size_out <= i{2} <= k{2} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ sp; if; auto=> />; 1,3:smt(domE rcons_cat). + move=> &1 &2 out_le_i _ ih1 ih2 ih3 i_lt_k xi_notin_mp r _. + by rewrite !get_set_sameE /= rcons_cat //= #smt:(get_setE size_out_gt0). +auto=> //=. +conseq(:_==> ={i0} /\ size bs{2} = i0{1} /\ (i0, x3){1} = (n, x2){2} /\ + bs0{1} = bs{2} /\ size bs{2} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ move=> />; smt(cats0 take_oversize spec_dout to_listK spec2_dout). +while(={i0} /\ x3{1} = x2{2} /\ 0 <= i0{1} <= n{2} /\ n{2} = size_out /\ + bs0{1} = bs{2} /\ size bs{2} = i0{1} /\ size_out <= n0{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ by sp; if; auto=> />; smt(size_rcons domE get_setE size_rcons mem_set). +by auto=> />; smt(size_out_gt0). +qed. + + +local lemma of_listK l : of_list (to_list l) = Some l. +proof. +by rewrite -to_listK. +qed. + +local module Fill_In (F : RO) = { + proc init = F.init + proc f (x : bool list, n : int) = { + var l, b, i; + i <- 0; + l <- []; + while (i < n) { + b <@ F.get((x,i)); + l <- rcons l b; + i <- i + 1; + } + while (i < size_out) { + F.sample((x,i)); + i <- i + 1; + } + return l; + } +}. + + +local equiv eq_eager_ideal : + BIRO.IRO.f ~ Fill_In(LRO).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp; rcondt{1} 1; auto. +while{2}(bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2})(size_out - i{2}). ++ by auto=> />; smt(). +conseq(:_==> bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}); 1: smt(). +while(={i, n, x} /\ bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). +by auto. +qed. + +local equiv eq_eager_ideal2 : + ExtendSample(FSome(BIRO.IRO)).f ~ FSome(Fill_In(RO)).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp. +if{1}; sp. ++ rcondt{1} 1; auto=> /=/>. + conseq(:_==> take k{1} bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). + * by move=> /> /#. + case: (0 <= n{2}); last first. + + rcondf{2} 1; 1: by auto; smt(). + conseq(:_==> BIRO.IRO.mp{1} = RO.m{2} /\ ={i} /\ n{1} = size_out /\ x2{1} = x0{2})=> />. + - smt(take_le0). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ BIRO.IRO.mp{1} = RO.m{2}). + - sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). + by auto=> />. + splitwhile{1} 1 : i < k. + while(={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + take k{1} bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} <= i{1} <= size_out). + * sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + by rcondf{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + conseq(:_==> ={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} = i{1}). + + smt(take_size). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ k{1} = n{2} /\ + 0 <= i{1} <= k{1} <= size_out /\ bs{1} = l{2} /\ size bs{1} = i{1} /\ + BIRO.IRO.mp{1} = RO.m{2}). + + sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(size_rcons). + by rcondf{2} 2; auto=> />; smt(size_rcons dbool_ll). + by auto=> />; smt(size_ge0 size_out_gt0). +rcondt{1} 1; auto. +rcondf{2} 2; 1: auto. ++ conseq(:_==> i = n)=> [/> /#|]. + by while(i <= n); auto=> />; smt(size_out_gt0). +while(i0{1} = i{2} /\ x3{1} = x0{2} /\ n0{1} = n{2} /\ bs0{1} = l{2} /\ + BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto; smt(dbool_ll). +by auto=> />. +qed. + +local module Dist (F : RO) = { + proc distinguish = SHA3_OIndiff.OIndif.OIndif(FSome(Fill_In(F)), + OSimulator(FSome(Fill_In(F))), ODRestr(Dist_of_P1Adv(A))).main +}. + +local module Game (F : RO) = { + proc distinguish () = { + var bo; + OSimulator(FSome(Fill_In(F))).init(); + Counter.c <- 0; + Log.m <- empty; + F.init(); + bo <@ Dist(F).distinguish(); + return bo; + } +}. + +local lemma eager_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res]. +proof. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] = + Pr[Game(LRO).distinguish() @ &m : res]. ++ byequiv=> //=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 2 2 : (={hash, m, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto. + call(: ={glob OFC, glob OSimulator} /\ BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:by auto=> /> &1 &2 <- /> <- />]; sim. + * inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto=> /> &1 &2 <- /> <- />. + by move=> /> &1 &2 <- /> <- />. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] = + Pr[Game(RO).distinguish() @ &m : res]. ++ byequiv=>//=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 2 2 : (={hash, m, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + by call eq_eager_ideal2; auto. + call(: ={glob OFC, glob OSimulator} /\ BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:by auto=> /> &1 &2 <- /> <- />]; sim. + * by call eq_eager_ideal2; auto=> /> &1 &2 <- /> <- />. + by move=> /> &1 &2 <- /> <- />. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + by call eq_eager_ideal2; auto. +rewrite eq_sym; byequiv=> //=; proc. +call(RO_LRO_D Dist _); first by rewrite dbool_ll. +by inline*; auto=> />. +qed. + +local lemma rw_ideal_2 &m: + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] <= + Pr[SORO.Preimage(SORO_P1(A), RFList).main() @ &m : res]. +proof. +have->:Pr[SORO.Preimage(SORO_P1(A), RFList).main() @ &m : res] = + Pr[SORO.Preimage(SORO_P1(A), RFWhile).main() @ &m : res]. ++ byequiv(: ={glob A} ==> _)=>//=; proc. + swap 1. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 2; inline{2} 2; sp. + swap[1..2] 3; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 5; inline{2} 5; wp. + seq 3 3 : (={mi, h, hash, glob A, glob SORO.Bounder, glob RFList}); last first. + - sp; if; auto; call(rw_RF_List_While); auto. + call(: ={glob SORO.Bounder, glob RFList, glob OSimulator, glob OPC, glob Log}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto; sim. + if; 1: auto; sim; sp; sim; if; auto=> [/> &1 &2 <- /> <- />||]; sim. + + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto=> />. + - by call(rw_RF_List_While); auto=> /> &1 &2 <- /> <- />. + by move=> /> &1 &2 <- /> <- />. + by move=> /> &1 &2 <- /> <- />. + - by sim. + proc; sim; inline{1} 1; inline{2} 1; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto; sim. + by call(rw_RF_List_While); auto. +rewrite (eager_ideal &m). +have->:Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendOutputSize(FSetSize(FSome(BIRO.IRO)))), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res]. ++ byequiv=> //=; proc; inline*; sp. + seq 2 2 : (={m, hash, glob OSimulator, glob OFC} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); last first. + - sp; if; auto; sp; if; auto; sp; rcondt{1}1; 1: auto. + * rcondt{2} 1; 1: auto. + while(={i, n, bs, x3} /\ size bs{1} = i{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} /\ + n{1} = size_out /\ 0 <= i{1} <= n{1}); auto. + * by sp; if; auto=> />; smt(domE get_setE size_rcons). + move=> />; smt(size_out_gt0 take_oversize size_out_gt0). + * by auto; rcondf{1} 1; auto. + * rcondt{2} 1; 1: auto; move=> />; auto. + by while(={i0, n0}); auto; sp; if{1}; if{2}; auto; smt(dbool_ll). + call(: ={glob OSimulator, glob OFC} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); last first; auto. + + smt(mem_empty). + + proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto. + if; 1, 3: auto; sp. + if; [1:auto=> /> &1 &2 <- /> <- />]; last first. + - by conseq=> />; sim=> /> &1 &2 <- /> <- />. + wp=> />; 1: smt(). + rnd; auto=> />. + call(eq_extend); last by auto=> /> &1 &2 <- /> <- /> /#. + + by proc; sp; if; auto; inline{1} 1; inline{2} 1; sp; if; auto. + proc; sp; inline{1} 1; inline{2} 1; sp; if; auto. + inline*; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp; auto. + conseq(:_==> ={bs} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); + 1: by auto. + while(={i, n, x3, bs} /\ 0 <= i{1} <= size_out /\ n{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + + sp; if; auto=> />. + + smt(domE get_setE size_rcons). + + move=> + + + + + + + + + + _. + smt(domE get_set_sameE get_setE size_rcons). + smt(domE get_setE size_rcons). + by auto=> />; smt(size_out_gt0). +byequiv=> //=; proc. +inline{1} 1; inline{2} 2; sp. +inline{1} 1; inline{2} 3; swap{2}[1..2]1; sp. +inline{1} 1; inline{2} 3; sp. +inline{1} 1; sp. +inline{1} 1; sp. +swap{2} 1 1; sp; swap{2}[1..2]3; sp. +inline{1} 1; sp; auto. +seq 2 5 : (={glob A, glob OSimulator, glob Counter, glob Log, hash, m} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); last first. ++ inline{1} 1; inline{2} 1; sp; inline{1} 1; sp; auto. + if{1}; sp; last first. + - conseq(:_==> true)=> />. + inline*; if{2}; auto; sp; if{2}; auto. + by while{2}(true)(size_out - i{2}); auto=>/>; smt(dbool_ll). + rcondt{2} 1; 1: by auto=> />; smt(divz_ge0 gt0_r size_ge0). + inline{1} 1; sp; auto. + rcondt{1} 1; auto=> /=. + inline{1} 1; sp; auto. + by call(eq_IRO_RFWhile); auto; smt(take_oversize). +auto; call(: ={glob OSimulator, glob Counter, glob Log} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); auto; last first. ++ by inline*; auto; smt(mem_empty). ++ proc; sp; if; auto=> />; 1: smt(). + inline{1} 1; inline{2} 1; sp; auto. + if; 1, 3: auto; -1: smt(). + if; 1,3:auto=> /> + + + + + + + + + + + _ + _ /#. + sp; if; [1:by auto=> /> &1 &2 <- /> <-]; last first. + + by conseq (: ={y, glob OSimulator}); [|sim]=> /> &1 &2 <- /> <- /#. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + rcondt{2} 1; 1: by auto; smt(). + sp. + seq 3 2 : (={x0, x1, o1, k0, Log.m, suffix, glob OSimulator} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{2} + 1); last first. + - by conseq(:_==> ={y, x1, glob OSimulator, Log.m}); 1: smt(); sim=> />. + inline{1} 1; auto. + by call(eq_IRO_RFWhile); auto=> /> &1 &2 + <- /> <- /#. ++ by proc; inline*; sp; if; auto; sp; if; auto=> />; smt(). +proc. +inline{1} 1; inline{2} 1; sp; if; auto=> /=. +inline{1} 1; inline{2} 1; sp. +rcondt{1} 1; 1: auto. +inline{1} 1; auto. +rcondf{2} 4; 1: auto. ++ inline*; auto; sp; if; auto; sp; if; auto=> />; conseq(:_==> true); 1: smt(). + by while(true); auto. +inline{2} 1; sp. +rcondt{2} 1; 1: by auto; smt(divz_ge0 gt0_r size_ge0). +auto; call eq_IRO_RFWhile; auto=> />. +move=> /> &l &r H0 H1 H2 H3 H5 H6 result_L mp_L m_R H7 H8 H9 H10 H11; split; 2: smt(divz_ge0 gt0_r size_ge0). +rewrite cats0 take_oversize 1:/# take_oversize 1:spec_dout //=. +have h:=spec2_dout result_L H7. +have-> := some_oget _ h. +by rewrite /= eq_sym -to_listK. +qed. + +local lemma rw_ideal &m: + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] <= + Pr[SORO.Preimage(SORO_P1(A),RF(SORO.RO.RO)).main() @ &m : res]. +proof. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal_2 &m)). +byequiv(: ={glob A} ==> _) => //=; proc; inline*; sp; wp. +swap{2} 2; sp; swap{2}[1..2] 6; sp. +swap{1} 2; sp; swap{1}[1..2] 6; sp. +seq 2 2 : ( + Log.m{1} = empty /\ + SHA3Indiff.Simulator.m{1} = empty /\ + SHA3Indiff.Simulator.mi{1} = empty /\ + SHA3Indiff.Simulator.paths{1} = empty.[c0 <- ([], b0)] /\ + Gconcl_list.BIRO2.IRO.mp{1} = empty /\ + SORO.Bounder.bounder{1} = 0 /\ + RFList.m{1} = empty /\ + Counter.c{2} = 0 /\ + ={Log.m, glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter} /\ + SORO.RO.RO.m{2} = empty /\ ={glob A, h, hash}); 1: auto=> />. +seq 1 1 : (={glob A, glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, + glob Log, mi, h, hash} /\ RFList.m{1} = SORO.RO.RO.m{2}). ++ call(: ={glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, glob Log} /\ + RFList.m{1} = SORO.RO.RO.m{2}); auto. + - proc; sp; if; 1, 3: auto; sp. + inline *; sp; sim. + if; 1: auto; sim. + if; 1: auto; sim. + sp; if; [2,3:sim]; [1,3:by auto=> /> &1 &2 <- /> <- />]. + sp; if; 1: auto; sim; -1: smt(). + sp; if{1}. + * rcondt{2} 2. + + by auto=> /> + + <- /> <- />. + auto. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> &l &r + <- /> <- /> - 7?; split; 1: smt(of_listK). + rewrite -dout_equal_dlist=> ?; split=> ?. + + by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite !get_setE/=dout_full/= => h. + rewrite eq_sym to_listK; apply some_oget. + apply spec2_dout. + by move:h; rewrite supp_dmap; smt(spec_dout). + by auto=> /> + + + <- /> <-; smt(dout_ll). + - by proc; inline*; sp; if; auto; sp; if; auto. + - proc; inline*; sp; if; auto; sp; if; auto; sp; sim. + if{1}. + * rcondt{2} 2; auto. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> &l *; split=> ?; 1: smt(of_listK). + rewrite -dout_equal_dlist; split=> ?. + * by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite supp_dmap dout_full/= =>/> a. + by rewrite get_setE/= dout_full/=; congr; rewrite of_listK oget_some. + by auto=> />; smt(dout_ll). +sp; if; 1, 3: auto; sp; wp 1 2. +if{1}. ++ wp=> />. + rnd (fun x => oget (of_list x)) to_list; auto=> />. + move=> &l c Hc Hnin; split. + - move=> ret Hret. + by have/= ->:= (to_listK ret (to_list ret)). + move=> h{h}; split. + - move=> ret Hret; rewrite -dout_equal_dlist. + rewrite dmapE /=; apply mu_eq=> //= x /=. + by rewrite /(\o) /pred1/=; smt(to_list_inj). + move=> h{h} l Hl. + rewrite dout_full /=. + have:= spec2_dout l. + have:=supp_dlist dbool size_out l _; 1: smt(size_out_gt0). + rewrite Hl/==> [#] -> h{h} /= H. + have H1:=some_oget _ H. + have:=to_listK (oget (of_list l)) l; rewrite {2}H1/= => -> /= {H H1}. + by rewrite get_setE/=. +by auto=> />; smt(dout_ll). +qed. + +local lemma leq_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res] <= (sigma + 1)%r / 2%r ^ size_out. +proof. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal &m)). +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (RO_is_preimage_resistant (SORO_P1(A)) &m)). +by rewrite doutE1. +qed. + + + + local lemma rw_real &m : + Pr[Preimage(A, OSponge, PSome(Perm)).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(Sponge(Poget(PSome(Perm)))), PSome(Perm), + ODRestr(Dist_of_P1Adv(A))).main() @ &m : res]. + proof. + byequiv=>//=; proc; inline*; sp; wp=> />. + swap{1} 4; sp. + seq 2 2 : (={glob A, glob Perm, hash, m} /\ Bounder.bounder{1} = Counter.c{2}). + + call(: ={glob Perm} /\ Bounder.bounder{1} = Counter.c{2})=> //=. +(** - by proc; inline*; sp; if; auto; 2:sim=> />; 1: smt(). **) +(** FIXME: two different instances of x{1} with InvalidGoalShape **) + - by proc; inline *; sp; if; auto; 2:sim=> />; smt(). + - by proc; inline*; sp; if; auto; 2:sim=> />; smt(). + - proc; inline*; sp; if; auto; sp=> />. + by conseq(:_==> ={z0, glob Perm})=> />; sim. + by auto. + by sp; if; auto=>/=; sim; auto. + qed. + +lemma Sponge_preimage_resistant &m: + (forall (F <: OIndif.ODFUNCTIONALITY) (P <: OIndif.ODPRIMITIVE), + islossless F.f => islossless P.f => islossless P.fi => islossless A(F,P).guess) => + Pr[Preimage(A, OSponge, PSome(Perm)).main() @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma + 1)%r / (2%r ^ size_out). +proof. +move=> A_ll. +rewrite (rw_real &m). +have := SHA3OIndiff (Dist_of_P1Adv(A)) &m _. ++ move=> F P Hp Hpi Hf; proc; inline*; sp; auto; call Hf; auto. + call(A_ll (DFSetSize(F)) P _ Hp Hpi); auto. + - proc; inline*; auto; call Hf; auto. + smt(dout_ll). +by have/#:=leq_ideal &m. +qed. + +end section Preimage. + + + +section SecondPreimage. + + + declare module A <: SH.AdvSecondPreimage {-Perm, -Counter, -Bounder, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -BIRO.IRO, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -SHA3Indiff.Simulator, -SHA3Indiff.Cntr, -SORO.Bounder, -SORO.RO.RO, -SORO.RO.FRO, -RO, -FRO}. + + local module FInit (F : OIndif.ODFUNCTIONALITY) : OIndif.OFUNCTIONALITY = { + proc init () = {} + proc f = F.f + }. + + local module PInit (P : ODPRIMITIVE) : OPRIMITIVE = { + proc init () = {} + proc f = P.f + proc fi = P.fi + }. + + +local module OF (F : Oracle) : OIndif.ODFUNCTIONALITY = { + proc f = F.get +}. + + +local module Log = { + var m : (bool list * int, bool) fmap +}. + +local module ExtendOutputSize (F : Oracle) : ODFUNCTIONALITY = { + proc f (x : bool list, k : int) = { + var o, l, suffix, prefix, i, r; + l <- None; + prefix <- []; + suffix <- []; + o <@ F.get(x); + prefix <- take k (to_list (oget o)); + i <- size_out; + while (i < k) { + if ((x,i) \notin Log.m) { + r <$ {0,1}; + Log.m.[(x,i)] <- r; + } + suffix <- rcons suffix (oget Log.m.[(x,i)]); + i <- i + 1; + } + l <- Some (prefix ++ suffix); + return l; + } +}. + +local module OFC2 (F : Oracle) = OFC(ExtendOutputSize(F)). + +local module ExtendOutput (F : RF) = { + proc init () = { + Log.m <- empty; + F.init(); + } + proc f = ExtendOutputSize(F).f + proc get = f +}. + + local module (Dist_of_P2Adv (A : SH.AdvSecondPreimage) : ODISTINGUISHER) (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + var m : bool list + proc distinguish () = { + var hash, hash', m'; + Log.m <- empty; + m' <@ A(DFSetSize(F),P).guess(m); + hash <@ DFSetSize(F).f(m); + hash' <@ DFSetSize(F).f(m'); + return m <> m' /\ exists y, hash' = Some y /\ hash = Some y; + } + }. + + +local module (SORO_P2 (A : SH.AdvSecondPreimage) : SORO.AdvSecondPreimage) (F : Oracle) = { + proc guess (m : bool list) : bool list = { + var mi; + Log.m <- empty; + Counter.c <- 0; + Dist_of_P2Adv.m <- m; + OSimulator(ExtendOutputSize(F)).init(); + mi <@ A(DFSetSize(OFC2(F)),OPC(OSimulator(ExtendOutputSize(F)))).guess(m); + return mi; + } +}. + +local module RFList = { + var m : (bool list, f_out) fmap + proc init () = { + m <- empty; + } + proc get (x : bool list) : f_out option = { + var z; + if (x \notin m) { + z <$ dlist dbool size_out; + m.[x] <- oget (of_list z); + } + return m.[x]; + } + proc sample (x: bool list) = {} +}. + +local module RFWhile = { + proc init () = { + RFList.m <- empty; + } + proc get (x : bool list) : f_out option = { + var l, i, b; + if (x \notin RFList.m) { + i <- 0; + l <- []; + while (i < size_out) { + b <$ dbool; + l <- rcons l b; + i <- i + 1; + } + RFList.m.[x] <- oget (of_list l); + } + return RFList.m.[x]; + } + proc sample (x: bool list) = {} +}. + + +local equiv rw_RF_List_While : + RFList.get ~ RFWhile.get : + ={arg, glob RFList} ==> ={res, glob RFWhile}. +proof. +proc; if; 1, 3: auto; wp. +conseq(:_==> z{1} = l{2})=> />. +transitivity{1} { + z <@ PBool.Sample.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>. ++ by inline*; auto. +transitivity{1} { + z <@ LoopSnoc.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>; last first. ++ inline*; auto; sim. + by while(={l, i} /\ n{1} = size_out); auto; smt(cats1). +by call(Sample_LoopSnoc_eq); auto. +qed. + + +local equiv eq_IRO_RFWhile : + BIRO.IRO.f ~ RFWhile.get : + arg{1} = (x{2}, size_out) /\ inv BIRO.IRO.mp{1} RFList.m{2} + ==> + res{2} = of_list res{1} /\ size res{1} = size_out /\ inv BIRO.IRO.mp{1} RFList.m{2}. +proof. +proc; inline*; sp. +rcondt{1} 1; 1: by auto. +if{2}; sp; last first. ++ alias{1} 1 mp = BIRO.IRO.mp. + conseq(:_==> BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ i{1} = size_out /\ + inv mp{1} RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))=> />. + - move=> &l &r H0 H1 H2 H3 H4 bs_L mp_L H5 H6 H7 H8 H9. + rewrite take_oversize 1:spec_dout 1:H5 //. + rewrite eq_sym to_listK => ->. + by have:=H4; rewrite domE; smt(). + - smt(take_oversize spec_dout). + while{1}(BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ + 0 <= i{1} <= size_out /\ n{1} = size_out /\ + inv mp{1} RFList.m{2} /\ x{1} \in RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))(size_out - i{1}); + auto=> />. + + sp; rcondf 1; auto=> />; 1: smt(). + move=> &h H0 H1 H2 H3 H4 H5 H6 H7 H8. + rewrite size_rcons //=; do!split; 1, 2, 4: smt(size_ge0). + rewrite (take_nth witness) 1:spec_dout 1:size_ge0//=. + rewrite - H7; congr; rewrite H5=> //=. + by apply H4=> //=. + smt(size_out_gt0 size_ge0 take0). +auto=> //=. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ n{1} = size_out /\ + inv BIRO.IRO.mp{1} RFList.m{2}.[x{2} <- oget (of_list l{2})])=> />. ++ smt(get_setE spec2_dout). ++ smt(get_setE spec2_dout). +alias{1} 1 m = BIRO.IRO.mp; sp. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ move=> /> &l &r H0 H1 H2 H3 H4 mp_L bs_L H5 H6 H7 H8 H9; do!split; ..-2 : smt(domE mem_set). + move=> l j Hin. + rewrite get_setE/=. + case: (l = x{r}) => [<<-|]. + - rewrite oget_some H9; 1:smt(); congr; congr. + by rewrite eq_sym to_listK; smt(spec2_dout). + move=> Hneq. + by rewrite -(H7 _ _ Hneq) H3; smt(domE). +while(l{2} = bs{1} /\ size bs{1} = i{1} /\ 0 <= i{1} <= n{1} /\ ={i} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ sp; rcondt{1} 1; auto=> />. + - smt(). + move=> &l &r H0 H1 H2 H3 H4 H5 H6 H7 H8 H9 H10 result_l _. + rewrite get_setE/=size_rcons/=; do!split; 1,2: smt(size_ge0). + - smt(mem_set). + - smt(get_setE). + - smt(mem_set). + - move=>j Hj0 Hjsize; rewrite get_setE/=nth_rcons. + case: (j = size bs{l})=>[->>//=|h]. + have/=Hjs:j < size bs{l} by smt(). + by rewrite Hjs/=H9//=. +by auto; smt(size_out_gt0). +qed. + + +local module ExtendSample (F : OFUNCTIONALITY) = { + proc init = F.init + proc f (x : bool list, k : int) = { + var y; + if (k <= size_out) { + y <@ F.f(x,size_out); + y <- omap (take k) y; + } else { + y <@ F.f(x,k); + } + return y; + } +}. + + +local equiv eq_extend : + ExtendSample(FSome(BIRO.IRO)).f ~ ExtendOutputSize(FSetSize(FSome(BIRO.IRO))).f : + ={arg} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} ==> + ={res} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}. +proof. +proc; inline*; auto; sp. +rcondt{2} 1; 1: auto. +if{1}; sp. +- rcondt{1} 1; auto. + rcondf{2} 8; 1: auto. + - conseq(:_==> true); 1: smt(). + by while(true); auto. + auto=> /=. + conseq(:_==> ={bs, k} /\ size bs{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2})=> //=. + - smt(cats0 to_listK spec2_dout). + while(={k, bs, n, x2} /\ i{1} = i0{2} /\ n{1} = size_out /\ + 0 <= i{1} <= n{1} /\ size bs{1} = i{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> />; smt(domE get_setE size_rcons). + by auto=> />; smt(size_eq0 size_out_gt0). +rcondt{1} 1; 1: auto. +splitwhile{1} 1 : i0 < size_out; auto=> /=. +while( (i0, n0, x3){1} = (i, k, x){2} /\ bs0{1} = prefix{2} ++ suffix{2} /\ + size_out <= i{2} <= k{2} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ sp; if; auto=> />. + + smt(). + + move=> + + + + + + + + + + _; smt(domE get_setE size_out_gt0 rcons_cat). + smt(domE get_setE size_out_gt0 rcons_cat). +auto=> //=. +conseq(:_==> ={i0} /\ size bs{2} = i0{1} /\ (i0, x3){1} = (n, x2){2} /\ + bs0{1} = bs{2} /\ size bs{2} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ move=> />; smt(cats0 take_oversize spec_dout to_listK spec2_dout). +while(={i0} /\ x3{1} = x2{2} /\ 0 <= i0{1} <= n{2} /\ n{2} = size_out /\ + bs0{1} = bs{2} /\ size bs{2} = i0{1} /\ size_out <= n0{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ by sp; if; auto=> />; smt(size_rcons domE get_setE size_rcons mem_set). +by auto=> />; smt(size_out_gt0). +qed. + + +local lemma of_listK l : of_list (to_list l) = Some l. +proof. +by rewrite -to_listK. +qed. + +local module Fill_In (F : RO) = { + proc init = F.init + proc f (x : bool list, n : int) = { + var l, b, i; + i <- 0; + l <- []; + while (i < n) { + b <@ F.get((x,i)); + l <- rcons l b; + i <- i + 1; + } + while (i < size_out) { + F.sample((x,i)); + i <- i + 1; + } + return l; + } +}. + + +local equiv eq_eager_ideal : + BIRO.IRO.f ~ Fill_In(LRO).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp; rcondt{1} 1; auto. +while{2}(bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2})(size_out - i{2}). ++ by auto=> />; smt(). +conseq(:_==> bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}); 1: smt(). +while(={i, n, x} /\ bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). +by auto. +qed. + +local equiv eq_eager_ideal2 : + ExtendSample(FSome(BIRO.IRO)).f ~ FSome(Fill_In(RO)).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp. +if{1}; sp. ++ rcondt{1} 1; auto=> /=/>. + conseq(:_==> take k{1} bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). + * by move=> />; smt(). + case: (0 <= n{2}); last first. + + rcondf{2} 1; 1: by auto; smt(). + conseq(:_==> BIRO.IRO.mp{1} = RO.m{2} /\ ={i} /\ n{1} = size_out /\ x2{1} = x0{2})=> />. + - smt(take_le0). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ BIRO.IRO.mp{1} = RO.m{2}). + - sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). + by auto=> />. + splitwhile{1} 1 : i < k. + while(={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + take k{1} bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} <= i{1} <= size_out). + * sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + by rcondf{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + conseq(:_==> ={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} = i{1}). + + by move=> />; smt(take_size). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ k{1} = n{2} /\ + 0 <= i{1} <= k{1} <= size_out /\ bs{1} = l{2} /\ size bs{1} = i{1} /\ + BIRO.IRO.mp{1} = RO.m{2}). + + sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(size_rcons). + by rcondf{2} 2; auto=> />; smt(size_rcons dbool_ll). + by auto=> />; smt(size_ge0 size_out_gt0). +rcondt{1} 1; auto. +rcondf{2} 2; 1: auto. ++ conseq(:_==> i = n); 1:by move=> />; smt(). + by while(i <= n); auto=> />; smt(size_out_gt0). +while(i0{1} = i{2} /\ x3{1} = x0{2} /\ n0{1} = n{2} /\ bs0{1} = l{2} /\ + BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto; smt(dbool_ll). +by auto=> />. +qed. + +local module Dist (F : RO) = { + proc distinguish = SHA3_OIndiff.OIndif.OIndif(FSome(Fill_In(F)), + OSimulator(FSome(Fill_In(F))), ODRestr(Dist_of_P2Adv(A))).main +}. + +local module Game (F : RO) = { + proc distinguish () = { + var bo; + OSimulator(FSome(Fill_In(F))).init(); + Counter.c <- 0; + Log.m <- empty; + F.init(); + bo <@ Dist(F).distinguish(); + return bo; + } +}. + +local lemma eager_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res]. +proof. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] = + Pr[Game(LRO).distinguish() @ &m : res]. ++ byequiv=> //=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 1 1 : (={m', glob Dist_of_P2Adv, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + * inline{1} 1; inline{2} 1; sp; sim. + inline{1} 7; inline{2} 7; sim. + inline{1} 8; inline{2} 8; sim. + swap 3 -2; sp. + case: (increase_counter Counter.c{1} m'{1} size_out <= SHA3Indiff.limit). + + rcondt{1} 10; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondt{2} 10; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + sim. + inline{1} 10; inline{2} 10; sim. + call eq_eager_ideal; auto. + by call eq_eager_ideal; auto. + rcondf{1} 10; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondf{2} 10; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + by auto; call eq_eager_ideal; auto. + sp; inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim. + if; auto. + inline{1} 1; inline{2} 1; sp; sim. + by auto; call eq_eager_ideal; auto. + call(: ={glob OFC, glob OSimulator, glob Dist_of_P2Adv} /\ + BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:auto; [1:by move=> /> + + <- /> <-]]; sim. + * inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto=> /> + + <- /> <- /#. + by move=> /> + + <- /> <- /#. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] = + Pr[Game(RO).distinguish() @ &m : res]. ++ byequiv=> //=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 1 1 : (={m', glob Dist_of_P2Adv, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + * inline{1} 6; inline{2} 6; sim. + inline{1} 7; inline{2} 7; sim. + swap 2 -1; sp. + case: (increase_counter Counter.c{1} m'{1} size_out <= SHA3Indiff.limit). + + rcondt{1} 9; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondt{2} 9; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + sim. + call eq_eager_ideal2; auto. + by call eq_eager_ideal2; auto. + rcondf{1} 9; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondf{2} 9; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + by auto; call eq_eager_ideal2; auto. + sp; inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim. + if; auto. + by auto; call eq_eager_ideal2; auto. + call(: ={glob OFC, glob OSimulator, glob Dist_of_P2Adv} /\ + BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:auto; [1:by move=> /> + + <- /> <-]]; sim. + * by call eq_eager_ideal2; auto=> /> + + <- /> <- /#. + by move=> /> + + <- /> <- /#. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + by call eq_eager_ideal2; auto. +rewrite eq_sym; byequiv=> //=; proc. +call(RO_LRO_D Dist _); first by rewrite dbool_ll. +by inline*; auto=> />. +qed. + + +local equiv toto : + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f ~ + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f : + ={glob OFC, arg} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} ==> + ={glob OFC, res} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}. +proof. +proc; inline*; sp; if; auto; sp; if; auto; sp; (rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto)=>/=. ++ conseq(:_==> ={bs} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); auto. + while(={i, bs, n, x3} /\ 0 <= i{1} <= size_out /\ n{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + + by sp; if; auto=> /> => [|+ + + + + + + + + + _|]; smt(domE get_setE size_out_gt0). + by auto=> />; smt(size_out_gt0). +by conseq(:_==> true); auto; sim. +qed. + +local equiv titi mess c: + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f + ~ + SORO.Bounder(RFWhile).get + : + ={arg} /\ arg{1} = mess /\ Counter.c{1} = c /\ + SORO.Bounder.bounder{2} <= Counter.c{1} /\ + inv BIRO.IRO.mp{1} RFList.m{2} + ==> + if (increase_counter c mess size_out <= sigma) then + (exists y, res{1} = Some y /\ res{2} = Some y /\ + SORO.Bounder.bounder{2} <= Counter.c{1} /\ + Counter.c{1} = increase_counter c mess size_out /\ + inv BIRO.IRO.mp{1} RFList.m{2}) + else (res{1} = None). +proof. +proc; sp. +inline{1} 1; sp; auto. +if{1}. +- rcondt{2} 1; first by auto; smt(divz_ge0 gt0_r size_ge0). + sp; auto. + inline{1} 1; sp; auto. + sp; rcondt{1} 1; auto. + inline{1} 1; sp; auto. + call(eq_IRO_RFWhile); auto=> /> &1 &2 bounder_R H0 H1 H2 H3 H4 H5 result_R mp_L m_R H6 H7 H8 H9. + rewrite take_oversize 1:/# /=. + have:=spec2_dout _ H6. + move=>/(some_oget)-> /=; smt(divz_ge0 gt0_r size_ge0 spec2_dout). +move=>/=. +conseq(:_==> true); auto. +inline*; if{2}; auto; sp; if{2}; auto; sp. +by while{2}(true)(size_out - i{2}); auto; smt(dbool_ll). +qed. + +local lemma rw_ideal_2 &m (mess : bool list): + Dist_of_P2Adv.m{m} = mess => + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] <= + Pr[SORO.SecondPreimage(SORO_P2(A), RFList).main(mess) @ &m : res]. +proof. +move=> Heq. +have->:Pr[SORO.SecondPreimage(SORO_P2(A), RFList).main(mess) @ &m : res] = + Pr[SORO.SecondPreimage(SORO_P2(A), RFWhile).main(mess) @ &m : res]. ++ byequiv(: ={glob A, arg} /\ arg{1} = mess ==> _)=>//=; proc. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + seq 1 1 : (={mi, m1, glob A, glob SORO.Bounder, glob RFList, glob Dist_of_P2Adv}); last first. + - sp; inline{1} 2; inline{2} 2; inline{1} 1; inline{2} 1; sp; sim. + if; auto. + - sp; case: (SORO.Bounder.bounder{1} < sigma). + * rcondt{1} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + rcondt{2} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + call(rw_RF_List_While); auto. + by call(rw_RF_List_While); auto=> />. + rcondf{1} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + rcondf{2} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + by auto; call(rw_RF_List_While); auto. + by sp; if; auto; call(rw_RF_List_While); auto. + call(: ={glob SORO.Bounder, glob RFList, glob OSimulator, glob OPC, glob Log, + glob Dist_of_P2Adv}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto; sim. + if; [1:by auto]; sim; sp; sim; if; auto=> /> => [+ + <- /> <- //||]; sim. + + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto=> />. + - by call(rw_RF_List_While); auto=> /> + + <- /> <-. + by move=> /> + + <- /> <-. + by move=> /> + + <- /> <-. + - by sim. + proc; sim; inline{1} 1; inline{2} 1; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto; sim. + by call(rw_RF_List_While); auto. +rewrite (eager_ideal &m). +have->:Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendOutputSize(FSetSize(FSome(BIRO.IRO)))), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res]. ++ byequiv=> //=; proc. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; auto=> />. + call(toto); call(toto); auto. + conseq(:_==> ={m', glob Counter, Dist_of_P2Adv.m} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); 1: smt(). + call(: ={glob OSimulator, glob OFC, Dist_of_P2Adv.m} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); last first; auto. + + smt(mem_empty). + + proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto. + if; 1, 3: auto; sp. + if; [1:by auto=> /> + + <- /> <-]; last first. + - by conseq=> />; sim=> /> + + <- /> <-. + wp=> />; [1:by auto=> /> + + <- /> <-]. + rnd; auto. + by call(eq_extend); auto=> /> + + <- /> <- /#. + + by proc; sp; if; auto; inline{1} 1; inline{2} 1; sp; if; auto. + proc; sp; inline{1} 1; inline{2} 1; sp; if; auto. + inline*; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp; auto. + conseq(:_==> ={bs} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); + 1: by auto. + while(={i, n, x3, bs} /\ 0 <= i{1} <= size_out /\ n{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> /> => [|+ + + + + + + + + + _|]; smt(domE get_setE size_rcons). + by auto=> />; smt(size_out_gt0). +byequiv=> //=; proc. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; sp; auto. +seq 1 1 : (={glob A, glob OFC, glob OSimulator, Log.m} /\ + m'{1} = mi{2} /\ m1{2} = Dist_of_P2Adv.m{1} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); last first. ++ sp; case: (increase_counter Counter.c{1} Dist_of_P2Adv.m{1} size_out <= SHA3Indiff.limit). + - exists * mi{2}, Dist_of_P2Adv.m{1}, Counter.c{1}; elim* => mess2 mess1 c. + call(titi mess2 (increase_counter c mess1 size_out))=> /=. + by call(titi mess1 c)=> />; auto; smt(). + inline*; sp. + rcondf{1} 1; 1: auto; sp. + conseq(:_==> true); auto. + seq 1 0 : true. + - if{1}; auto; sp; 1: if{1}; auto; sp. + - rcondt{1} 1; auto. + while{1}(true)(n1{1}-i1{1}); auto; -1: smt(). + by sp; if; auto; smt(dbool_ll). + rcondt{1} 1; 1: auto. + while{1}(true)(n2{1}-i2{1}); auto. + by sp; if; auto; smt(dbool_ll). + seq 0 1 : true. + - if{2}; auto; sp; if{2}; auto; sp. + by while{2}(true)(size_out-i{2}); auto; smt(dbool_ll). + sp; if{2}; auto; sp; if{2}; auto; sp. + by while{2}(true)(size_out-i0{2}); auto; smt(dbool_ll). +conseq(:_==> ={glob A, glob OFC, glob OSimulator, Log.m} /\ + m'{1} = mi{2} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ SORO.Bounder.bounder{2} <= Counter.c{1}). ++ smt(). +auto; call(: ={glob OSimulator, glob Counter, glob Log} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); auto; last first. ++ by smt(mem_empty). ++ proc; sp; if; auto=> />; 1: smt(). + inline{1} 1; inline{2} 1; sp; auto. + if; 1, 3: auto; -1: smt(). + if; [1,3:auto]; [2:move=> /> + + + + + + + + + + + _ + _ - /#]. + sp; if; [1:by auto=> /> + + <- /> <-]; last first. + - conseq(:_==> ={y, glob OSimulator}). + + by auto=> /> + + <- /> <- /#. + by sim=> /> + + <- /> <-. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + rcondt{2} 1; 1: by auto; smt(). + sp. + seq 3 2 : (={x0, x1, o1, k0, Log.m, suffix, glob OSimulator} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{2} + 1); last first. + - by conseq(:_==> ={y, x1, glob OSimulator, Log.m}); 1: smt(); sim=> />. + inline{1} 1; auto. + by call(eq_IRO_RFWhile); auto=> /> + + + <- /> <- /#. ++ by proc; inline*; sp; if; auto; sp; if; auto=> />; smt(). +proc. +inline{1} 1; inline{2} 1; sp; if; auto=> /=. +inline{1} 1; inline{2} 1; sp. +rcondt{1} 1; 1: auto. +inline{1} 1; auto. +rcondf{2} 4; 1: auto. ++ inline*; auto; sp; if; auto; sp; if; auto=> />; conseq(:_==> true); 1: smt(). + by while(true); auto. +inline{2} 1; sp. +rcondt{2} 1; 1: by auto; smt(divz_ge0 gt0_r size_ge0). +auto; call eq_IRO_RFWhile; auto=> />. +move=> &l &r H0 H1 H2 H3 H4 H5 result_L mp_L m_R H6 H7 H8 H9 H10; split; 2: smt(divz_ge0 gt0_r size_ge0). +rewrite cats0 take_oversize 1:/# take_oversize 1:spec_dout //=. +have h:=spec2_dout result_L H6. +have-> := some_oget _ h. +by rewrite eq_sym -to_listK; congr. +qed. + +local lemma rw_ideal &m (mess : bool list): + Dist_of_P2Adv.m{m} = mess => + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] <= + Pr[SORO.SecondPreimage(SORO_P2(A),RF(SORO.RO.RO)).main(mess) @ &m : res]. +proof. +move=> Heq. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal_2 &m mess Heq)). +byequiv(: ={glob A} /\ ={arg} /\ arg{1} = mess ==> _) => //=; proc; inline*; sp; wp. +seq 1 1 : (={glob A, glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, + glob Log, mi, m1} /\ RFList.m{1} = SORO.RO.RO.m{2}). ++ call(: ={glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, glob Log} /\ + RFList.m{1} = SORO.RO.RO.m{2}); auto. + - proc; sp; if; 1, 3: auto; sp. + inline *; sp; sim. + if; 1: auto; sim. + if; 1: auto; sim. + sp; if; [1:by auto=> /> + + <- /> <-]; sim; 2:by auto=> /> + + <- /> <-. + sp; if; 1: auto; sim; -1: smt(). + sp; if{1}. + * rcondt{2} 2; auto; [1:by auto=> /> + + <- /> <-; smt(BlockSponge.parse_valid)]. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> /> &l &r + <- /> <- /> - 7?; split; 1: smt(of_listK). + rewrite -dout_equal_dlist=> ?; split=> ?. + + by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite !get_setE/=dout_full/= => h. + rewrite eq_sym to_listK; apply some_oget. + apply spec2_dout. + by move:h; rewrite supp_dmap; smt(spec_dout). + by auto=> /> + + + <- /> <-; smt(dout_ll). + - by proc; inline*; sp; if; auto; sp; if; auto. + - proc; inline*; sp; if; auto; sp; if; auto; sp; sim. + if{1}. + * rcondt{2} 2; auto. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> &l 4?; split=> ?; 1: smt(of_listK). + rewrite -dout_equal_dlist; split=> ?. + * by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite supp_dmap dout_full/= =>/> a. + by rewrite get_setE/= dout_full/=; congr; rewrite of_listK oget_some. + by auto=> />; smt(dout_ll). +sp. +seq 4 4 : (={SORO.Bounder.bounder, x0, m1, m2, hash1, y0} /\ y0{1} = None /\ + RFList.m{1} = SORO.RO.RO.m{2}); last first. ++ if; 1, 3: auto; sp. + if{1}. + - rcondt{2} 2; 1: auto. + auto; rnd (fun t => oget (of_list t)) to_list; auto=> />. + move=> &l c Hc Hnin; split. + - move=> ret Hret. + by have/= ->:= (to_listK ret (to_list ret)). + move=> h{h}; split. + - move=> ret Hret; rewrite -dout_equal_dlist. + rewrite dmapE /=; apply mu_eq=> //= x /=. + by rewrite /(\o) /pred1/=; smt(to_list_inj). + move=> h{h} l Hl. + rewrite dout_full /=. + have:= spec2_dout l. + have:=supp_dlist dbool size_out l _; 1: smt(size_out_gt0). + rewrite Hl/==> [#] -> h{h} /= H. + have H1:=some_oget _ H. + have:=to_listK (oget (of_list l)) l; rewrite {2}H1/= => -> /= {H H1}. + by rewrite get_setE/=; smt(). + by auto=> />; smt(dout_ll). +if; 1, 3: auto; sp. +if{1}. +- rcondt{2} 2; 1: auto. + auto; rnd (fun t => oget (of_list t)) to_list; auto=> />. + move=> &l c Hc Hnin; split. + - move=> ret Hret. + by have/= ->:= (to_listK ret (to_list ret)). + move=> h{h}; split. + - move=> ret Hret; rewrite -dout_equal_dlist. + rewrite dmapE /=; apply mu_eq=> //= x /=. + by rewrite /(\o) /pred1/=; smt(to_list_inj). + move=> h{h} l Hl. + rewrite dout_full /=. + have:= spec2_dout l. + have:=supp_dlist dbool size_out l _; 1: smt(size_out_gt0). + rewrite Hl/==> [#] -> h{h} /= H. + have H1:=some_oget _ H. + have:=to_listK (oget (of_list l)) l; rewrite {2}H1/= => -> /= {H H1}. + by rewrite get_setE/=; smt(). +by auto=> />; smt(dout_ll). +qed. + + +local lemma leq_ideal &m mess: + Dist_of_P2Adv.m{m} = mess => + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res] <= (sigma + 1)%r / 2%r ^ size_out. +proof. +move=> Heq. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal &m mess Heq)). +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (RO_is_second_preimage_resistant (SORO_P2(A)) &m mess)). +by rewrite doutE1. +qed. + +local lemma rw_real &m mess : + Dist_of_P2Adv.m{m} = mess => + Pr[SecondPreimage(A, OSponge, PSome(Perm)).main(mess) @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(Sponge(Poget(PSome(Perm)))), PSome(Perm), + ODRestr(Dist_of_P2Adv(A))).main() @ &m : res]. +proof. +move=> Heq. +byequiv=>//=; proc. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; sp; wp=> />. +seq 1 1 : (={glob A, glob Perm} /\ m1{1} = Dist_of_P2Adv.m{2} /\ + m2{1} = m'{2} /\ Bounder.bounder{1} = Counter.c{2}). ++ auto; call(: ={glob Perm} /\ Bounder.bounder{1} = Counter.c{2})=> //=. + - by proc; inline*; sp; if; auto; 2:sim=> />; smt(). + - by proc; inline*; sp; if; auto; 2:sim=> />; smt(). + - proc; inline*; sp; if; auto; sp=> />. + by conseq(:_==> ={z0, glob Perm})=> />; sim. + by auto; smt(). +conseq(:_==> m1{1} = Dist_of_P2Adv.m{2} /\ m2{1} = m'{2} /\ + hash1{1} = hash{2} /\ hash2{1} = hash'{2})=> //=; 1: smt(). +seq 1 1 : (m1{1} = Dist_of_P2Adv.m{2} /\ m2{1} = m'{2} /\ + hash1{1} = hash{2} /\ ={glob Perm} /\ Bounder.bounder{1} = Counter.c{2}); last first. ++ inline*; sp; if; auto; sp=> /=; sim. +inline*; sp; if; auto; swap{1} 9; auto; sp=> /=. +by conseq(:_==> m1{1} = Dist_of_P2Adv.m{2} /\ m2{1} = m'{2} /\ + of_list (oget (Some (take n{1} z0{1}))) = + of_list (oget (Some (take n{2} z0{2}))) /\ ={Perm.mi, Perm.m})=> //=; sim. +qed. + +local module TOTO = { + proc main (m : bool list) = { + var b; + Dist_of_P2Adv.m <- m; + b <@ SecondPreimage(A, OSponge, PSome(Perm)).main(m); + return b; + } +}. + +lemma Sponge_second_preimage_resistant &m mess: + (forall (F <: OIndif.ODFUNCTIONALITY) (P <: OIndif.ODPRIMITIVE), + islossless F.f => islossless P.f => islossless P.fi => islossless A(F,P).guess) => + Pr[SecondPreimage(A, OSponge, PSome(Perm)).main(mess) @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma + 1)%r / (2%r ^ size_out). +proof. +move=> A_ll. +have->:Pr[SecondPreimage(A, OSponge, PSome(Perm)).main(mess) @ &m : res] = + Pr[TOTO.main(mess) @ &m : res]. ++ by byequiv=> //=; proc; inline*; auto; sim. +byphoare(: arg = mess ==>_)=>//=; proc; sp. +call(: arg = mess /\ mess = Dist_of_P2Adv.m ==> res); auto. +bypr=> {&m} &m [#]->; rewrite eq_sym=> Heq. +rewrite (rw_real &m mess Heq). +have := SHA3OIndiff (Dist_of_P2Adv(A)) &m _. ++ move=> F P Hp Hpi Hf; proc; inline*; sp; auto; call Hf; auto; call Hf; auto. + call(A_ll (DFSetSize(F)) P _ Hp Hpi); auto. + proc; inline*; auto; call Hf; auto. +by have/#:=leq_ideal &m. +qed. + +end section SecondPreimage. + + + + +section Collision. + + + declare module A <: SH.AdvCollision {-Perm, -Counter, -Bounder, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -BIRO.IRO, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -SHA3Indiff.Simulator, -SHA3Indiff.Cntr, -SORO.Bounder, -SORO.RO.RO, -SORO.RO.FRO, -RO, -FRO}. + + local module FInit (F : OIndif.ODFUNCTIONALITY) : OIndif.OFUNCTIONALITY = { + proc init () = {} + proc f = F.f + }. + + local module PInit (P : ODPRIMITIVE) : OPRIMITIVE = { + proc init () = {} + proc f = P.f + proc fi = P.fi + }. + + +local module OF (F : Oracle) : OIndif.ODFUNCTIONALITY = { + proc f = F.get +}. + + +local module Log = { + var m : (bool list * int, bool) fmap +}. + +local module ExtendOutputSize (F : Oracle) : ODFUNCTIONALITY = { + proc f (x : bool list, k : int) = { + var o, l, suffix, prefix, i, r; + l <- None; + prefix <- []; + suffix <- []; + o <@ F.get(x); + prefix <- take k (to_list (oget o)); + i <- size_out; + while (i < k) { + if ((x,i) \notin Log.m) { + r <$ {0,1}; + Log.m.[(x,i)] <- r; + } + suffix <- rcons suffix (oget Log.m.[(x,i)]); + i <- i + 1; + } + l <- Some (prefix ++ suffix); + return l; + } +}. + +local module OFC2 (F : Oracle) = OFC(ExtendOutputSize(F)). + +local module ExtendOutput (F : RF) = { + proc init () = { + Log.m <- empty; + F.init(); + } + proc f = ExtendOutputSize(F).f + proc get = f +}. + + local module (Dist_of_CollAdv (A : SH.AdvCollision) : ODISTINGUISHER) (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + var m : bool list + proc distinguish () = { + var hash1, hash2, m1, m2; + Log.m <- empty; + (m1, m2) <@ A(DFSetSize(F),P).guess(); + hash1 <@ DFSetSize(F).f(m1); + hash2 <@ DFSetSize(F).f(m2); + return m1 <> m2 /\ exists y, hash1 = Some y /\ hash2 = Some y; + } + }. + + +local module (SORO_Coll (A : SH.AdvCollision) : SORO.AdvCollision) (F : Oracle) = { + proc guess () = { + var mi; + Log.m <- empty; + Counter.c <- 0; + OSimulator(ExtendOutputSize(F)).init(); + mi <@ A(DFSetSize(OFC2(F)),OPC(OSimulator(ExtendOutputSize(F)))).guess(); + return mi; + } +}. + +local module RFList = { + var m : (bool list, f_out) fmap + proc init () = { + m <- empty; + } + proc get (x : bool list) : f_out option = { + var z; + if (x \notin m) { + z <$ dlist dbool size_out; + m.[x] <- oget (of_list z); + } + return m.[x]; + } + proc sample (x: bool list) = {} +}. + +local module RFWhile = { + proc init () = { + RFList.m <- empty; + } + proc get (x : bool list) : f_out option = { + var l, i, b; + if (x \notin RFList.m) { + i <- 0; + l <- []; + while (i < size_out) { + b <$ dbool; + l <- rcons l b; + i <- i + 1; + } + RFList.m.[x] <- oget (of_list l); + } + return RFList.m.[x]; + } + proc sample (x: bool list) = {} +}. + + +local equiv rw_RF_List_While : + RFList.get ~ RFWhile.get : + ={arg, glob RFList} ==> ={res, glob RFWhile}. +proof. +proc; if; 1, 3: auto; wp. +conseq(:_==> z{1} = l{2})=> />. +transitivity{1} { + z <@ PBool.Sample.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>. ++ by inline*; auto. +transitivity{1} { + z <@ LoopSnoc.sample(size_out); + } + (true ==> ={z}) + (true ==> z{1} = l{2})=>/>; last first. ++ inline*; auto; sim. + by while(={l, i} /\ n{1} = size_out); auto; smt(cats1). +by call(Sample_LoopSnoc_eq); auto. +qed. + + +local equiv eq_IRO_RFWhile : + BIRO.IRO.f ~ RFWhile.get : + arg{1} = (x{2}, size_out) /\ inv BIRO.IRO.mp{1} RFList.m{2} + ==> + res{2} = of_list res{1} /\ size res{1} = size_out /\ inv BIRO.IRO.mp{1} RFList.m{2}. +proof. +proc; inline*; sp. +rcondt{1} 1; 1: by auto. +if{2}; sp; last first. ++ alias{1} 1 mp = BIRO.IRO.mp. + conseq(:_==> BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ i{1} = size_out /\ + inv mp{1} RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))=> />. + - move=> &l &r H0 H1 H2 H3 H4 bs_L mp_L H5 H7 H8 H9 H10. + rewrite take_oversize 1:spec_dout 1:H5 //. + rewrite eq_sym to_listK => ->. + by have:=H4; rewrite domE; smt(). + - smt(take_oversize spec_dout). + while{1}(BIRO.IRO.mp{1} = mp{1} /\ size bs{1} = i{1} /\ + 0 <= i{1} <= size_out /\ n{1} = size_out /\ + inv mp{1} RFList.m{2} /\ x{1} \in RFList.m{2} /\ + bs{1} = take i{1} (to_list (oget RFList.m{2}.[x{1}])))(size_out - i{1}); + auto=> />. + + sp; rcondf 1; auto=> />; 1: smt(). + move=> &h H0 H1 H2 H3 H4 H5 H6 H7 H8. + rewrite size_rcons //=; do!split; 1, 2, 4: smt(size_ge0). + rewrite (take_nth witness) 1:spec_dout 1:size_ge0//=. + rewrite - H7; congr; rewrite H5=> //=. + by apply H4=> //=. + smt(size_out_gt0 size_ge0 take0). +auto=> //=. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ n{1} = size_out /\ + inv BIRO.IRO.mp{1} RFList.m{2}.[x{2} <- oget (of_list l{2})])=> />. ++ smt(get_setE spec2_dout). ++ smt(get_setE spec2_dout). +alias{1} 1 m = BIRO.IRO.mp; sp. +conseq(:_==> l{2} = bs{1} /\ size bs{1} = i{1} /\ i{1} = n{1} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ move=> /> &l &r H0 H1 H2 H3 H4 mp_L bs_L H5 H6 H7 H8 H9; do!split; ..-2 : smt(domE mem_set). + move=> l j Hin. + rewrite get_setE/=. + case: (l = x{r}) => [<<-|]. + - rewrite oget_some H9; 1:smt(); congr; congr. + by rewrite eq_sym to_listK; smt(spec2_dout). + move=> Hneq. + by rewrite -(H7 _ _ Hneq) H3; smt(domE). +while(l{2} = bs{1} /\ size bs{1} = i{1} /\ 0 <= i{1} <= n{1} /\ ={i} /\ + n{1} = size_out /\ inv m{1} RFList.m{2} /\ + (forall j, (x{1}, j) \in BIRO.IRO.mp{1} => 0 <= j < i{1}) /\ + (forall l j, l <> x{1} => m{1}.[(l,j)] = BIRO.IRO.mp{1}.[(l,j)]) /\ + (forall j, 0 <= j < i{1} => (x{1}, j) \in BIRO.IRO.mp{1}) /\ + (forall j, 0 <= j < i{1} => BIRO.IRO.mp{1}.[(x{1},j)] = Some (nth witness bs{1} j))). ++ sp; rcondt{1} 1; auto=> />. + - smt(). + move=> &l &r H0 H1 H2 H3 H4 H5 H6 H7 H8 H9 H10 rL _. + rewrite get_setE/=size_rcons/=; do!split; 1,2: smt(size_ge0). + - smt(mem_set). + - smt(get_setE). + - smt(mem_set). + - move=>j Hj0 Hjsize; rewrite get_setE/=nth_rcons. + case: (j = size bs{l})=>[->>//=|h]. + have/=Hjs:j < size bs{l} by smt(). + by rewrite Hjs/=H9//=. +by auto; smt(size_out_gt0). +qed. + + +local module ExtendSample (F : OFUNCTIONALITY) = { + proc init = F.init + proc f (x : bool list, k : int) = { + var y; + if (k <= size_out) { + y <@ F.f(x,size_out); + y <- omap (take k) y; + } else { + y <@ F.f(x,k); + } + return y; + } +}. + + +local equiv eq_extend : + ExtendSample(FSome(BIRO.IRO)).f ~ ExtendOutputSize(FSetSize(FSome(BIRO.IRO))).f : + ={arg} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} ==> + ={res} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}. +proof. +proc; inline*; auto; sp. +rcondt{2} 1; 1: auto. +if{1}; sp. +- rcondt{1} 1; auto. + rcondf{2} 8; 1: auto. + - conseq(:_==> true); 1: smt(). + by while(true); auto. + auto=> /=. + conseq(:_==> ={bs, k} /\ size bs{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2})=> //=. + - smt(cats0 to_listK spec2_dout). + while(={k, bs, n, x2} /\ i{1} = i0{2} /\ n{1} = size_out /\ + 0 <= i{1} <= n{1} /\ size bs{1} = i{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> />; smt(domE get_setE size_rcons). + by auto=> />; smt(size_eq0 size_out_gt0). +rcondt{1} 1; 1: auto. +splitwhile{1} 1 : i0 < size_out; auto=> /=. +while( (i0, n0, x3){1} = (i, k, x){2} /\ bs0{1} = prefix{2} ++ suffix{2} /\ + size_out <= i{2} <= k{2} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ by sp; if; auto=> /> => [|+ + + + + + + + + + _|]; smt(domE get_setE size_out_gt0 rcons_cat). +auto=> //=. +conseq(:_==> ={i0} /\ size bs{2} = i0{1} /\ (i0, x3){1} = (n, x2){2} /\ + bs0{1} = bs{2} /\ size bs{2} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ by auto=> />; smt(cats0 take_oversize spec_dout to_listK spec2_dout). +while(={i0} /\ x3{1} = x2{2} /\ 0 <= i0{1} <= n{2} /\ n{2} = size_out /\ + bs0{1} = bs{2} /\ size bs{2} = i0{1} /\ size_out <= n0{1} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). ++ by sp; if; auto=> /> => [|+ + + + + + + + + + + + _|]; smt(size_rcons domE get_setE size_rcons mem_set). +by auto=> />; smt(size_out_gt0). +qed. + + +local lemma of_listK l : of_list (to_list l) = Some l. +proof. +by rewrite -to_listK. +qed. + +local module Fill_In (F : RO) = { + proc init = F.init + proc f (x : bool list, n : int) = { + var l, b, i; + i <- 0; + l <- []; + while (i < n) { + b <@ F.get((x,i)); + l <- rcons l b; + i <- i + 1; + } + while (i < size_out) { + F.sample((x,i)); + i <- i + 1; + } + return l; + } +}. + + +local equiv eq_eager_ideal : + BIRO.IRO.f ~ Fill_In(LRO).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp; rcondt{1} 1; auto. +while{2}(bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2})(size_out - i{2}). ++ by auto=> />; smt(). +conseq(:_==> bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}); 1: smt(). +while(={i, n, x} /\ bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). +by auto. +qed. + +local equiv eq_eager_ideal2 : + ExtendSample(FSome(BIRO.IRO)).f ~ FSome(Fill_In(RO)).f : + ={arg} /\ BIRO.IRO.mp{1} = RO.m{2} ==> + ={res} /\ BIRO.IRO.mp{1} = RO.m{2}. +proof. +proc; inline*; sp. +if{1}; sp. ++ rcondt{1} 1; auto=> /=/>. + conseq(:_==> take k{1} bs{1} = l{2} /\ BIRO.IRO.mp{1} = RO.m{2}). + * by auto=> />; smt(). + case: (0 <= n{2}); last first. + + rcondf{2} 1; 1: by auto; smt(). + conseq(:_==> BIRO.IRO.mp{1} = RO.m{2} /\ ={i} /\ n{1} = size_out /\ x2{1} = x0{2})=> />. + - smt(take_le0). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ BIRO.IRO.mp{1} = RO.m{2}). + - sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto=> />; smt(dbool_ll). + by auto=> />. + splitwhile{1} 1 : i < k. + while(={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + take k{1} bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} <= i{1} <= size_out). + * sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + by rcondf{2} 2; auto=> />; smt(dbool_ll cats1 take_cat cats0 take_size size_rcons). + conseq(:_==> ={i} /\ n{1} = size_out /\ x2{1} = x0{2} /\ BIRO.IRO.mp{1} = RO.m{2} /\ + bs{1} = l{2} /\ size bs{1} = i{1} /\ k{1} = i{1}). + + smt(take_size). + while(={i} /\ x2{1} = x0{2} /\ n{1} = size_out /\ k{1} = n{2} /\ + 0 <= i{1} <= k{1} <= size_out /\ bs{1} = l{2} /\ size bs{1} = i{1} /\ + BIRO.IRO.mp{1} = RO.m{2}). + + sp; if{1}. + - by rcondt{2} 2; auto=> />; smt(size_rcons). + by rcondf{2} 2; auto=> />; smt(size_rcons dbool_ll). + by auto=> />; smt(size_ge0 size_out_gt0). +rcondt{1} 1; auto. +rcondf{2} 2; 1: auto. ++ conseq(:_==> i = n)=> />; 1:smt(). + by while(i <= n); auto=> />; smt(size_out_gt0). +while(i0{1} = i{2} /\ x3{1} = x0{2} /\ n0{1} = n{2} /\ bs0{1} = l{2} /\ + BIRO.IRO.mp{1} = RO.m{2}). ++ sp; if{1}. + - by rcondt{2} 2; auto=> />. + by rcondf{2} 2; auto; smt(dbool_ll). +by auto=> />. +qed. + +local module Dist (F : RO) = { + proc distinguish = SHA3_OIndiff.OIndif.OIndif(FSome(Fill_In(F)), + OSimulator(FSome(Fill_In(F))), ODRestr(Dist_of_CollAdv(A))).main +}. + +local module Game (F : RO) = { + proc distinguish () = { + var bo; + OSimulator(FSome(Fill_In(F))).init(); + Counter.c <- 0; + Log.m <- empty; + F.init(); + bo <@ Dist(F).distinguish(); + return bo; + } +}. + +local lemma eager_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res]. +proof. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), + OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] = + Pr[Game(LRO).distinguish() @ &m : res]. ++ byequiv=> //=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 1 1 : (={m1, m2, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + * inline{1} 1; inline{2} 1; sp; sim. + inline{1} 7; inline{2} 7; sim. + inline{1} 8; inline{2} 8; sim. + swap 3 -2; sp. + case: (increase_counter Counter.c{1} m2{1} size_out <= SHA3Indiff.limit). + + rcondt{1} 10; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondt{2} 10; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + sim. + inline{1} 10; inline{2} 10; sim. + call eq_eager_ideal; auto. + by call eq_eager_ideal; auto. + rcondf{1} 10; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondf{2} 10; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + by auto; call eq_eager_ideal; auto. + sp; inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim. + if; auto. + inline{1} 1; inline{2} 1; sp; sim. + by auto; call eq_eager_ideal; auto. + call(: ={glob OFC, glob OSimulator} /\ + BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:by auto=> /> + + <- /> <-]; sim. + * inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto=> /> + + <- /> <-. + by auto=> /> + + <- /> <-. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + inline{1} 1; inline{2} 1; sp; sim. + by call eq_eager_ideal; auto. +have->: + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] = + Pr[Game(RO).distinguish() @ &m : res]. ++ byequiv=> //=; proc. + inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; sim. + seq 1 1 : (={m1, m2, glob OFC} /\ BIRO.IRO.mp{1} = RO.m{2}); last first. + - inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim; if; auto. + * inline{1} 6; inline{2} 6; sim. + inline{1} 7; inline{2} 7; sim. + swap 2 -1; sp. + case: (increase_counter Counter.c{1} m2{1} size_out <= SHA3Indiff.limit). + + rcondt{1} 9; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondt{2} 9; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + sim. + call eq_eager_ideal2; auto. + by call eq_eager_ideal2; auto. + rcondf{1} 9; 1: auto. + - inline*; auto. + by sp; rcondt 1; auto; conseq(:_==> true); auto. + rcondf{2} 9; 1: auto. + - inline*; auto. + by conseq(:_==> true); auto. + by auto; call eq_eager_ideal2; auto. + sp; inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; sim. + if; auto. + by auto; call eq_eager_ideal2; auto. + call(: ={glob OFC, glob OSimulator} /\ + BIRO.IRO.mp{1} = RO.m{2}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + if; 1: auto; sim; sp. + if; [1:by auto=> /> + + <- /> <-]; sim. + * by call eq_eager_ideal2; auto=> /> + + <- /> <-. + by auto=> /> + + <- /> <-. + - by proc; inline*; sim. + proc; sim. + inline{1} 1; inline{2} 1; sp; sim; if; 1: auto; sim. + by call eq_eager_ideal2; auto. +rewrite eq_sym; byequiv=> //=; proc. +call(RO_LRO_D Dist _); first by rewrite dbool_ll. +by inline*; auto=> />. +qed. + +local equiv toto : + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f ~ + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f : + ={glob OFC, arg} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2} ==> + ={glob OFC, res} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}. +proof. +proc; inline*; sp; if; auto; sp; if; auto; sp; (rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto)=>/=. ++ conseq(:_==> ={bs} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); auto. + while(={i, bs, n, x3} /\ 0 <= i{1} <= size_out /\ n{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> /> => [|+ + + + + + + + + + _|]; smt(domE get_setE size_out_gt0). + by auto=> />; smt(size_out_gt0). +by conseq(:_==> true); auto; sim. +qed. + +local equiv titi mess c: + DFSetSize(OFC(ExtendSample(FSome(BIRO.IRO)))).f + ~ + SORO.Bounder(RFWhile).get + : + ={arg} /\ arg{1} = mess /\ Counter.c{1} = c /\ + SORO.Bounder.bounder{2} <= Counter.c{1} /\ + inv BIRO.IRO.mp{1} RFList.m{2} + ==> + if (increase_counter c mess size_out <= sigma) then + (exists y, res{1} = Some y /\ res{2} = Some y /\ + SORO.Bounder.bounder{2} <= Counter.c{1} /\ + Counter.c{1} = increase_counter c mess size_out /\ + inv BIRO.IRO.mp{1} RFList.m{2}) + else (res{1} = None). +proof. +proc; sp. +inline{1} 1; sp; auto. +if{1}. +- rcondt{2} 1; first by auto; smt(divz_ge0 gt0_r size_ge0). + sp; auto. + inline{1} 1; sp; auto. + sp; rcondt{1} 1; auto. + inline{1} 1; sp; auto. + call(eq_IRO_RFWhile); auto=> /> &1 &2 bounder_R H0 H1 H2 H3 H4 H5 result_L mp_L m_R H6 H7 H8 H9. + rewrite take_oversize 1:/# /=. + have:=spec2_dout _ H6. + move=>/(some_oget)-> /=; smt(divz_ge0 gt0_r size_ge0 spec2_dout). +move=>/=. +conseq(:_==> true); auto. +inline*; if{2}; auto; sp; if{2}; auto; sp. +by while{2}(true)(size_out - i{2}); auto; smt(dbool_ll). +qed. + +local lemma rw_ideal_2 &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] <= + Pr[SORO.Collision(SORO_Coll(A), RFList).main() @ &m : res]. +proof. +have->:Pr[SORO.Collision(SORO_Coll(A), RFList).main() @ &m : res] = + Pr[SORO.Collision(SORO_Coll(A), RFWhile).main() @ &m : res]. ++ byequiv(: ={glob A, arg} ==> _)=>//=; proc. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + seq 1 1 : (={mi, glob A, glob SORO.Bounder, glob RFList}); last first. + - sp; inline{1} 2; inline{2} 2; inline{1} 1; inline{2} 1; sp; sim. + if; auto. + - sp; case: (SORO.Bounder.bounder{1} < sigma). + * rcondt{1} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + rcondt{2} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + call(rw_RF_List_While); auto. + by call(rw_RF_List_While); auto=> />. + rcondf{1} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + rcondf{2} 5; 1: auto. + + by inline*; auto; conseq(:_==> true); auto. + by auto; call(rw_RF_List_While); auto. + by sp; if; auto; call(rw_RF_List_While); auto. + call(: ={glob SORO.Bounder, glob RFList, glob OSimulator, glob OPC, glob Log}); auto. + - proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto; sim. + if; [1:by auto]; sim; sp; sim; if; auto=> /> => [+ + <- /> <- //||]; sim. + + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto=> />. + - by call(rw_RF_List_While); auto=> /> + + <- /> <-. + by auto=> /> + + <- /> <-. + by auto=> /> + + <- /> <-. + - by sim. + proc; sim; inline{1} 1; inline{2} 1; sp; if; auto. + inline{1} 1; inline{2} 1; sp; sim. + inline{1} 1; inline{2} 1; sp; if; auto; sim. + by call(rw_RF_List_While); auto. +rewrite (eager_ideal &m). +have->:Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendSample(FSome(BIRO.IRO))), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(ExtendSample(FSome(BIRO.IRO)), + OSimulator(ExtendOutputSize(FSetSize(FSome(BIRO.IRO)))), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res]. ++ byequiv=> //=; proc. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp; auto=> />. + call(toto); call(toto); auto. + conseq(:_==> ={m1, m2, glob Counter} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); 1: smt(). + call(: ={glob OSimulator, glob OFC} /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); last first; auto. + + smt(mem_empty). + + proc; sp; if; auto. + inline{1} 1; inline{2} 1; sp; if; 1, 3: auto. + if; 1, 3: auto; sp. + if; [1:by auto=> /> + + <- /> <-]; last first. + - by conseq=> />; sim=> /> + + <- /> <-. + wp=> />; 1: smt(). + rnd; auto. + by call(eq_extend); auto=> /> + + <- /> <- /#. + by proc; sp; if; auto; inline{1} 1; inline{2} 1; sp; if; auto. + proc; sp; inline{1} 1; inline{2} 1; sp; if; auto. + inline*; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp. + rcondt{1} 1; 1: auto; rcondt{2} 1; 1: auto; sp; auto. + conseq(:_==> ={bs} /\ eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}); + 1: by auto. + while(={i, n, x3, bs} /\ 0 <= i{1} <= size_out /\ n{1} = size_out /\ + eq_extend_size BIRO.IRO.mp{1} BIRO.IRO.mp{2} Log.m{2}). + - by sp; if; auto=> /> => [|+ + + + + + + + + + _|]; smt(domE get_setE size_rcons). + by auto=> />; smt(size_out_gt0). +byequiv=> //=; proc. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; sp; auto. +seq 1 2 : (={glob A, glob OFC, glob OSimulator, Log.m, m1, m2} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); last first. ++ sp; case: (increase_counter Counter.c{1} m1{1} size_out <= SHA3Indiff.limit). + - exists * m2{2}, m1{1}, Counter.c{1}; elim* => mess2 mess1 c. + call(titi mess2 (increase_counter c mess1 size_out))=> /=. + by call(titi mess1 c)=> />; auto; smt(). + inline*; sp. + rcondf{1} 1; 1: auto; sp. + conseq(:_==> true); auto. + seq 1 0 : true. + - if{1}; auto; sp; 1: if{1}; auto; sp. + - rcondt{1} 1; auto. + while{1}(true)(n1{1}-i1{1}); auto; -1: smt(). + by sp; if; auto; smt(dbool_ll). + rcondt{1} 1; 1: auto. + while{1}(true)(n2{1}-i2{1}); auto. + by sp; if; auto; smt(dbool_ll). + seq 0 1 : true. + - if{2}; auto; sp; if{2}; auto; sp. + by while{2}(true)(size_out-i{2}); auto; smt(dbool_ll). + sp; if{2}; auto; sp; if{2}; auto; sp. + by while{2}(true)(size_out-i0{2}); auto; smt(dbool_ll). +conseq(:_==> ={glob A, glob OFC, glob OSimulator, Log.m, m1, m2} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ SORO.Bounder.bounder{2} <= Counter.c{1}). +auto; call(: ={glob OSimulator, glob Counter, glob Log} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{1}); auto; last first. ++ by smt(mem_empty). ++ proc; sp; if; auto=> />; 1: smt(). + inline{1} 1; inline{2} 1; sp; auto. + if; 1, 3: auto; -1: smt(). + if; [1,3:auto]; 2:by move=> /> + + + + + + + + + + + _ + _ /#. + sp; if; [1:by auto=> /> + + <- /> <-]; last first. + - by conseq(:_==> ={y, glob OSimulator}); [2:sim]; auto=> /> + + <- /> <- /#. + inline{1} 1; inline{2} 1; sp. + inline{1} 1; inline{2} 1; sp. + rcondt{2} 1; 1: by auto; smt(). + sp. + seq 3 2 : (={x0, x1, o1, k0, Log.m, suffix, glob OSimulator} /\ + inv BIRO.IRO.mp{1} RFList.m{2} /\ + SORO.Bounder.bounder{2} <= Counter.c{2} + 1); last first. + - by conseq(:_==> ={y, x1, glob OSimulator, Log.m}); 1: smt(); sim=> />. + inline{1} 1; auto. + by call(eq_IRO_RFWhile); auto=> /> + + + <- /> <- /#. ++ by proc; inline*; sp; if; auto; sp; if; auto=> />; smt(). +proc. +inline{1} 1; inline{2} 1; sp; if; auto=> /=. +inline{1} 1; inline{2} 1; sp. +rcondt{1} 1; 1: auto. +inline{1} 1; auto. +rcondf{2} 4; 1: auto. ++ inline*; auto; sp; if; auto; sp; if; auto=> />; conseq(:_==> true); 1: smt(). + by while(true); auto. +inline{2} 1; sp. +rcondt{2} 1; 1: by auto; smt(divz_ge0 gt0_r size_ge0). +auto; call eq_IRO_RFWhile; auto=> />. +move=> &l &r H0 H1 H2 H3 H4 H5 result_L mp_L m_R H6 H7 H8 H9 H10; split; 2: smt(divz_ge0 gt0_r size_ge0). +rewrite cats0 take_oversize 1:/# take_oversize 1:spec_dout //=. +have h:=spec2_dout result_L H6. +have-> := some_oget _ h. +by rewrite eq_sym -to_listK; congr. +qed. + +local lemma rw_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] <= + Pr[SORO.Collision(SORO_Coll(A),RF(SORO.RO.RO)).main() @ &m : res]. +proof. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal_2 &m)). +byequiv(: ={glob A} ==> _) => //=; proc; inline*; sp; wp. +seq 1 1 : (={glob A, glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, + glob Log, mi} /\ RFList.m{1} = SORO.RO.RO.m{2}). ++ call(: ={glob SHA3Indiff.Simulator, glob SORO.Bounder, glob Counter, glob Log} /\ + RFList.m{1} = SORO.RO.RO.m{2}); auto. + - proc; sp; if; 1, 3: auto; sp. + inline *; sp; sim. + if; 1: auto; sim. + if; 1: auto; sim. + sp; if; [3:sim]; [1,3:by auto=> /> + + <- /> <-]. + sp; if; 1: auto; sim; -1: smt(). + sp; if{1}. + * rcondt{2} 2; auto; [1:by auto=> /> + + <- /> <-; smt(BlockSponge.parse_valid)]. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> /> &l &r + <- /> <- /> 6?; split; 1: smt(of_listK). + rewrite -dout_equal_dlist=> ?; split=> ?. + + by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite !get_setE/= dout_full/= => h. + rewrite eq_sym to_listK; apply some_oget. + apply spec2_dout. + by move:h; rewrite supp_dmap; smt(spec_dout). + by auto=> /> + + + <- /> <-; smt(dout_ll). + - by proc; inline*; sp; if; auto; sp; if; auto. + - proc; inline*; sp; if; auto; sp; if; auto; sp; sim. + if{1}. + * rcondt{2} 2; auto. + rnd (fun l => oget (of_list l)) to_list; auto=> />. + move=> &l 4?; split=> ?; 1: smt(of_listK). + rewrite -dout_equal_dlist; split=> ?. + * by rewrite dmapE=> h{h}; apply mu_eq=> x; smt(to_list_inj). + move=> sample. + rewrite supp_dmap dout_full/= =>/> a. + by rewrite get_setE/= dout_full/=; congr; rewrite of_listK oget_some. + by auto=> />; smt(dout_ll). +sp. +seq 4 4 : (={SORO.Bounder.bounder, x0, m1, m2, hash1, y0} /\ y0{1} = None /\ + RFList.m{1} = SORO.RO.RO.m{2}); last first. ++ if; 1, 3: auto; sp. + if{1}. + - rcondt{2} 2; 1: auto. + auto; rnd (fun t => oget (of_list t)) to_list; auto=> />. + move=> &l c Hc Hnin; split. + - move=> ret Hret. + by have/= ->:= (to_listK ret (to_list ret)). + move=> h{h}; split. + - move=> ret Hret; rewrite -dout_equal_dlist. + rewrite dmapE /=; apply mu_eq=> //= x /=. + by rewrite /(\o) /pred1/=; smt(to_list_inj). + move=> h{h} l Hl. + rewrite dout_full /=. + have:= spec2_dout l. + have:=supp_dlist dbool size_out l _; 1: smt(size_out_gt0). + rewrite Hl/==> [#] -> h{h} /= H. + have H1:=some_oget _ H. + have:=to_listK (oget (of_list l)) l; rewrite {2}H1/= => -> /= {H H1}. + by rewrite get_setE/=; smt(). + by auto=> />; smt(dout_ll). +if; 1, 3: auto; sp. +if{1}. +- rcondt{2} 2; 1: auto. + auto; rnd (fun t => oget (of_list t)) to_list; auto=> />. + move=> &l c Hc Hnin; split. + - move=> ret Hret. + by have/= ->:= (to_listK ret (to_list ret)). + move=> h{h}; split. + - move=> ret Hret; rewrite -dout_equal_dlist. + rewrite dmapE /=; apply mu_eq=> //= x /=. + by rewrite /(\o) /pred1/=; smt(to_list_inj). + move=> h{h} l Hl. + rewrite dout_full /=. + have:= spec2_dout l. + have:=supp_dlist dbool size_out l _; 1: smt(size_out_gt0). + rewrite Hl/==> [#] -> h{h} /= H. + have H1:=some_oget _ H. + have:=to_listK (oget (of_list l)) l; rewrite {2}H1/= => -> /= {H H1}. + by rewrite get_setE/=; smt(). +by auto=> />; smt(dout_ll). +qed. + +local lemma leq_ideal &m : + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(BIRO.IRO), OSimulator(FSome(BIRO.IRO)), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res] <= + (sigma * (sigma - 1) + 2)%r / 2%r / 2%r ^ size_out. +proof. +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (rw_ideal &m)). +rewrite (StdOrder.RealOrder.ler_trans _ _ _ (RO_is_collision_resistant (SORO_Coll(A)) &m)). +by rewrite doutE1. +qed. + +local lemma rw_real &m : + Pr[Collision(A, OSponge, PSome(Perm)).main() @ &m : res] = + Pr[SHA3_OIndiff.OIndif.OIndif(FSome(Sponge(Poget(PSome(Perm)))), PSome(Perm), + ODRestr(Dist_of_CollAdv(A))).main() @ &m : res]. +proof. +byequiv=>//=; proc. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; inline{2} 1; sp. +inline{1} 1; sp; wp=> />. +seq 1 1 : (={glob A, glob Perm, m1, m2} /\ Bounder.bounder{1} = Counter.c{2}). ++ auto; call(: ={glob Perm} /\ Bounder.bounder{1} = Counter.c{2})=> //=. + - by proc; inline*; sp; if; auto; 2:sim=> />; smt(). + - by proc; inline*; sp; if; auto; 2:sim=> />; smt(). + - proc; inline*; sp; if; auto; sp=> />. + by conseq(:_==> ={z0, glob Perm})=> />; sim. +conseq(:_==> ={hash1, hash2, m1, m2})=> //=; 1: smt(); sim. +seq 1 1 : (={m1, m2, hash1, glob Perm} /\ Bounder.bounder{1} = Counter.c{2}); last first. ++ inline*; sp; if; auto; sp=> /=; sim. +inline*; sp; if; auto; swap{1} 9; auto; sp=> /=. +by conseq(:_==> ={m1, m2} /\ of_list (oget (Some (take n{1} z0{1}))) = + of_list (oget (Some (take n{2} z0{2}))) /\ ={Perm.mi, Perm.m})=> //=; sim. +qed. + +lemma Sponge_collision_resistant &m : + (forall (F <: OIndif.ODFUNCTIONALITY) (P <: OIndif.ODPRIMITIVE), + islossless F.f => islossless P.f => islossless P.fi => islossless A(F,P).guess) => + Pr[Collision(A, OSponge, PSome(Perm)).main() @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma * (sigma - 1) + 2)%r / 2%r / 2%r ^ size_out. +proof. +move=> A_ll. +rewrite (rw_real &m). +have := SHA3OIndiff (Dist_of_CollAdv(A)) &m _. ++ move=> F P Hp Hpi Hf; proc; inline*; sp; auto; call Hf; auto; call Hf; auto. + call(A_ll (DFSetSize(F)) P _ Hp Hpi); auto. + proc; inline*; auto; call Hf; auto. +by have/#:=leq_ideal &m. +qed. + +end section Collision. diff --git a/sha3/proof/SHA3Security.ec b/sha3/proof/SHA3Security.ec new file mode 100644 index 0000000..6fe3370 --- /dev/null +++ b/sha3/proof/SHA3Security.ec @@ -0,0 +1,1274 @@ +(* Top-level Proof of SHA-3 Security *) + +require import AllCore Distr DList DBool List IntDiv Dexcepted DProd SmtMap FSet. +require import Common SLCommon Sponge SHA3Indiff. +(*---*) import StdOrder.IntOrder. +require (****) IndifRO_is_secure. + +module SHA3 (P : DPRIMITIVE) = { + proc init() : unit = {} + proc f (bl : bool list, n : int) : bool list = { + var r : bool list; + r <@ Sponge(P).f(bl ++ [false; true], n); + return r; + } +}. + +op size_out : int. +axiom size_out_gt0 : 0 < size_out. + +op sigma : int. +axiom sigma_gt0 : 0 < sigma. + +type f_out. + +op dout : f_out distr. +axiom dout_ll : is_lossless dout. +axiom dout_fu : is_funiform dout. +axiom dout_full : is_full dout. + + +op to_list : f_out -> bool list. +op of_list : bool list -> f_out option. +axiom spec_dout (l : f_out) : size (to_list l) = size_out. +axiom spec2_dout (l : bool list) : size l = size_out => of_list l <> None. +axiom to_list_inj : injective to_list. +axiom to_listK e l : to_list e = l <=> of_list l = Some e. + +axiom dout_equal_dlist : dmap dout to_list = dlist dbool size_out. + +lemma doutE1 x : mu1 dout x = inv (2%r ^ size_out). +proof. +have->:inv (2%r ^ size_out) = mu1 (dlist dbool size_out) (to_list x). ++ rewrite dlist1E. + - smt(size_out_gt0). + rewrite spec_dout/=. + pose p:= StdBigop.Bigreal.BRM.big _ _ _. + have->: p = StdBigop.Bigreal.BRM.big predT (fun _ => inv 2%r) (to_list x). + - rewrite /p =>{p}. + apply StdBigop.Bigreal.BRM.eq_bigr. + by move=> i; rewrite//= dbool1E. + rewrite StdBigop.Bigreal.BRM.big_const count_predT spec_dout=> {p}. + have:=size_out_gt0; move/ltzW. + move:size_out;apply intind=> //=. + - by rewrite RField.expr0 iter0 //= fromint1. + move=> i hi0 rec. + by rewrite RField.exprS //iterS// -rec; smt(). +rewrite -dout_equal_dlist dmap1E. +apply mu_eq. +by move=> l; rewrite /pred1/(\o); smt(to_listK). +qed. + +module CSetSize (F : CONSTRUCTION) (P : DPRIMITIVE) = { + proc init = F(P).init + proc f (x : bool list) = { + var r; + r <@ F(P).f(x,size_out); + return oget (of_list r); + } +}. + +module FSetSize (F : FUNCTIONALITY) = { + proc init = F.init + proc f (x : bool list) = { + var r; + r <@ F.f(x,size_out); + return oget (of_list r); + } +}. + +clone import IndifRO_is_secure as S with + type block <- block * capacity, + type f_in <- bool list, + type f_out <- f_out, + + op sampleto <- dout, + op bound <- (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r, + op limit <- sigma, + op bound_counter <- limit, + op increase_counter <- fun c m => c + ((size m + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 + + proof *. + + +realize bound_counter_ge0 by exact(SLCommon.max_ge0). +realize limit_gt0 by exact(sigma_gt0). +realize sampleto_ll by exact(dout_ll). +realize sampleto_fu by exact(dout_fu). +realize sampleto_full by exact(dout_full). +realize increase_counter_spec by smt(List.size_ge0 divz_ge0 gt0_r). + +module FGetSize (F : Indiff0.DFUNCTIONALITY) = { + proc f (x : bool list, i : int) = { + var r; + r <@ F.f(x); + return to_list r; + } +}. + +module SimSetSize (S : SIMULATOR) (F : Indiff0.DFUNCTIONALITY) = S(FGetSize(F)). + +module DFSetSize (F : DFUNCTIONALITY) = { + proc f (x : bool list) = { + var r; + r <@ F.f(x,size_out); + return oget (of_list r); + } +}. + +module (DSetSize (D : Indiff0.DISTINGUISHER) : DISTINGUISHER) + (F : DFUNCTIONALITY) (P : DPRIMITIVE) = D(DFSetSize(F),P). + + +section Preimage. + + declare module A <: SRO.AdvPreimage {-SRO.RO.RO, -SRO.RO.FRO, -SRO.Bounder, -Perm, -Gconcl_list.BIRO2.IRO, -Simulator, -Cntr, -BIRO.IRO, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -DPre}. + + declare axiom A_ll (F <: SRO.Oracle {-A}) : islossless F.get => islossless A(F).guess. + + local lemma invm_dom_rng (m mi : (state, state) fmap) : + invm m mi => dom m = rng mi. + proof. + move=>h; rewrite fun_ext=> x; rewrite domE rngE /= eq_iff; have h2 := h x; split. + + move=> m_x_not_None; exists (oget m.[x]); rewrite -h2; move: m_x_not_None. + by case: (m.[x]). + by move=> [] a; rewrite -h2 => ->. + qed. + + local lemma invmC' (m mi : (state, state) fmap) : + invm m mi => invm mi m. + proof. by rewrite /#. qed. + + local lemma invmC (m mi : (state, state) fmap) : + invm m mi <=> invm mi m. + proof. by split;exact invmC'. qed. + + local lemma useful m mi a : + Prefix.invm m mi => ! a \in m => Distr.is_lossless ((bdistr `*` cdistr) \ rng m). + proof. + move=>hinvm nin_dom. + have prod_ll:Distr.is_lossless (bdistr `*` cdistr). + + by rewrite dprod_ll DBlock.dunifin_ll DCapacity.dunifin_ll. + apply dexcepted_ll=>//=;rewrite-prod_ll. + have->:predT = predU (predC (rng m)) (rng m);1:rewrite predCU//=. + rewrite Distr.mu_disjoint 1:predCI//=RField.addrC. + have/=->:=StdOrder.RealOrder.ltr_add2l (mu (bdistr `*` cdistr) (rng m)) 0%r. + rewrite Distr.witness_support/predC. + move:nin_dom;apply absurd=>//=;rewrite negb_exists/==>hyp. + have{hyp}hyp:forall x, rng m x by smt(supp_dprod DBlock.supp_dunifin DCapacity.supp_dunifin). + move:a. + have:=eqEcard (fdom m) (frng m);rewrite leq_card_rng_dom/=. + have->//=:fdom m \subset frng m. + + by move=> x; rewrite mem_fdom mem_frng hyp. + smt(mem_fdom mem_frng). + qed. + + + local equiv equiv_sponge_perm c m : + FInit(CSetSize(Sponge, Perm)).get ~ FInit(DFSetSize(FC(Sponge(Perm)))).get : + ={arg, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c /\ arg{2} = m /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} ==> + ={res, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline FC(Sponge(Perm)).f; sp. + rcondt{2} 1; auto; sp. + call(: ={glob Perm} /\ invm Perm.m{1} Perm.mi{1})=>/=; auto; inline*. + while(={i, n, sa, sc, z, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + + sp; if; auto; sp; if; auto; progress. + rewrite invm_set //=. + by move:H4; rewrite supp_dexcepted. + sp; conseq(:_==> ={i, n, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + while(={xs, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + sp; if; auto; progress. + rewrite invm_set=>//=. + by move:H4; rewrite supp_dexcepted. + qed. + + + op same_ro (m1 : (bool list, f_out) fmap) (m2 : (bool list * int, bool) fmap) = + (forall m, m \in m1 => forall i, 0 <= i < size_out => (m,i) \in m2) + && (forall m, (exists i, 0 <= i < size_out /\ (m,i) \in m2) => m \in m1) + && (forall m, m \in m1 => to_list (oget m1.[m]) = map (fun i => oget m2.[(m,i)]) (range 0 size_out)). + + op same_ro2 (m1 : (bool list, bool list) fmap) (m2 : (bool list * int, bool) fmap) = + (forall m, m \in m1 => forall i, 0 <= i < size_out => (m,i) \in m2) + && (forall m, (exists i, 0 <= i < size_out /\ (m,i) \in m2) => m \in m1) + && (forall m, m \in m1 => oget m1.[m] = map (fun i => oget m2.[(m,i)]) (range 0 size_out)). + + clone import Program as Prog with + type t <- bool, + op d <- dbool + proof *. + + local equiv equiv_ro_iro c m : + FInit(RO).get ~ FInit(DFSetSize(FC(BIRO.IRO))).get : + ={arg} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + arg{2} = m /\ Cntr.c{2} = c /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} + ==> ={res} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline *; sp; rcondt{2} 1; 1: auto. + swap{2} 1 5; sp; wp 2 1. + conseq(:_==> oget SRO.RO.RO.m{1}.[x{1}] = oget (of_list bs0{2}) /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); 1:by auto. + rcondt{2} 1; 1: auto. + case: (x{1} \in SRO.RO.RO.m{1}). + + rcondf{1} 2; auto. + exists* BIRO.IRO.mp{2}; elim* => mp. + while{2}(bs0{2} = map (fun j => oget BIRO.IRO.mp{2}.[(x0{2},j)]) (range 0 i{2}) + /\ n0{2} = size_out /\ x0{2} \in SRO.RO.RO.m{1} /\ 0 <= i{2} <= size_out + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ BIRO.IRO.mp{2} = mp) + (size_out - i{2}); auto. + - sp; rcondf 1; auto; 1: smt(). + progress. + * have/=<-:= map_rcons (fun (j : int) => oget BIRO.IRO.mp{hr}.[(x0{hr}, j)]) (range 0 i{hr}) i{hr}. + by rewrite !rangeSr //=. + * smt(). + * smt(). + * smt(). + progress. + - by rewrite range_geq. + - smt(size_out_gt0). + - smt(). + - exact(dout_ll). + - have[] h[#] h1 h2 := H. + have->:i_R = size_out by smt(). + have<-:=h2 _ H3. + smt(to_listK). + rcondt{1} 2; 1: auto; wp =>/=. + exists* BIRO.IRO.mp{2}; elim* => mp. + conseq(:_==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress=>//=. + + by rewrite get_set_sameE /=; smt(to_listK take_oversize spec_dout). + + move:H8; rewrite mem_set=>[][]//=h; 1:rewrite H3=>//=. + - by have []h1 []h2 h3:= H2; have->//:=h1 _ h. + by move:h => <<-; rewrite H6 //=. + + rewrite mem_set //=; have [] //= h:= H5 _ _ H11; left. + have []h1 []->//=:= H2. + by exists i0=>//=. + + move:H7; rewrite take_oversize 1:spec_dout//= => H7. + move:H10; rewrite mem_set. + case(m \in SRO.RO.RO.m{1})=>//=h. + - rewrite get_set_neqE; 1:smt(). + have []h1 []h2 ->//=:= H2. + by apply eq_in_map=> j;rewrite mem_range=>[][]hj1 hj2/=; rewrite H4//=h1//=. + by move=><<-; rewrite get_set_eqE//=. + alias{1} 1 l = [<:bool>]. + transitivity {1} { + l <@ Sample.sample(size_out); + r <- oget (of_list l); + } + (={glob SRO.RO.RO, x} ==> ={glob SRO.RO.RO, r}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x{1} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + inline*; sp; wp. + rnd to_list (fun x => oget (of_list x)); auto; progress. + - smt(spec_dout supp_dlist to_listK spec2_dout size_out_gt0). + - rewrite -dout_equal_dlist dmap1E; apply mu_eq=> x/=. + smt(to_listK). + - rewrite-dout_equal_dlist supp_dmap; smt(dout_full). + smt(to_listK). + wp=>/=. + conseq(:_==> i{2} = size_out /\ size l{1} = size_out /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => (l0, j) \in BIRO.IRO.mp{2}) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => BIRO.IRO.mp{2}.[(l0, j)] = mp.[(l0, j)]) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in BIRO.IRO.mp{2} => ((l0, j) \in mp) \/ (l0 = x0{2} /\ 0 <= j < i{2})) /\ + (forall (j : int), 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + transitivity{1} { + l <@ LoopSnoc.sample(size_out); + } + (={glob SRO.RO.RO} ==> ={glob SRO.RO.RO, l}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x0{2} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + i{2} = size_out /\ size l{1} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + by call Sample_LoopSnoc_eq; auto. + inline*; sp; wp. + conseq(:_==> i{2} = size_out /\ size l0{1} = i{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress. + + smt(take_oversize). + + smt(take_oversize). + while(0 <= i{2} <= size_out /\ size l0{1} = i{2} /\ n0{2} = size_out /\ + ={i} /\ n{1} = n0{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})). + + sp; wp=> //=. + rcondt{2} 1; 1:auto; progress. + - have[]h1 [] h2 h3 := H1. + have:=h2 x0{hr}; rewrite H2/= negb_exists/= =>/(_ (size bs0{hr})). + rewrite size_ge0 H9/=; apply absurd =>/= h. + by have //=:= H5 _ _ h. + wp; rnd; auto; progress. + - smt(size_ge0). + - smt(). + - by rewrite size_cat/=. + - by rewrite mem_set; left; rewrite H3. + - rewrite get_setE (H4 _ _ H11). + have/#: !(l1, j) = (x0{2}, size bs0{2}). + move:H2; apply absurd=> //=[#] <<- ->>. + have[] h1 [] h2 h3 := H1. + by apply h2; smt(). + - move:H11; rewrite mem_set. + case((l1, j) \in BIRO.IRO.mp{2})=>//= h; 1: smt(). + by move=> [#] <<- ->> //=; rewrite size_ge0; smt(). + - rewrite mem_set. + case(j = size bs0{2})=>//=. + move=> h; rewrite h /=; have {H12} H13 {h} : j < size bs0{2} by smt(). + by apply H6. + - by rewrite cats1 get_set_sameE oget_some. + - rewrite get_set_sameE oget_some H7 rangeSr. + rewrite !size_map 1:size_ge0. + rewrite (size_map _ (range 0 (size bs0{2}))) size_range /=. + rewrite ler_maxr 1:size_ge0 map_rcons /=get_set_sameE oget_some; congr. + apply eq_in_map=> j. + rewrite mem_range /==> [] [] hj1 hj2. + by rewrite get_set_neqE //=; smt(). + auto; progress. + + smt(size_out_gt0). + + smt(). + + smt(). + + by rewrite range_geq. + smt(). + qed. + + lemma Sponge_preimage_resistant &m ha : + (DPre.h{m} = ha) => + Pr[SRO.Preimage(A, FM(CSetSize(Sponge), Perm)).main(ha) @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma + 1)%r / (2%r ^ size_out). + proof. + move=>init_ha. + rewrite -(doutE1 ha). + rewrite(preimage_resistant_if_indifferentiable A A_ll (CSetSize(Sponge)) Perm &m ha init_ha). + exists (SimSetSize(Simulator))=>//=; split. + + by move=> F _; proc; inline*; auto. + have->//:Pr[Indiff0.Indif(CSetSize(Sponge, Perm), Perm, DPre(A)).main() @ &m : res] = + Pr[RealIndif(Sponge, Perm, DRestr(DSetSize(DPre(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline DPre(A, CSetSize(Sponge, Perm), Perm).distinguish. + inline SRO.Preimage(A, FInit(CSetSize(Sponge, Perm))).main. + inline DRestr(DSetSize(DPre(A)), Sponge(Perm), Perm).distinguish + DSetSize(DPre(A), FC(Sponge(Perm)), PC(Perm)).distinguish + SRO.Preimage(A, FInit(DFSetSize(FC(Sponge(Perm))))).main. + inline Perm.init CSetSize(Sponge, Perm).init Sponge(Perm).init + FC(Sponge(Perm)).init SRO.Counter.init Cntr.init + SRO.Bounder(FInit(CSetSize(Sponge, Perm))).init + SRO.Bounder(FInit(DFSetSize(FC(Sponge(Perm))))).init + FInit(CSetSize(Sponge, Perm)).init + FInit(DFSetSize(FC(Sponge(Perm)))).init; sp. + wp; sp; sim. + seq 1 1 : (={m, hash, glob DPre, glob SRO.Counter, glob Perm} + /\ invm Perm.m{1} Perm.mi{1} /\ DPre.h{1} = ha + /\ ={c}(SRO.Counter,Cntr)); last first. + - if; auto; sp. + exists* m{1}, SRO.Counter.c{1}; elim* => mess c. + by call(equiv_sponge_perm c mess); auto; smt(). + call(: ={glob SRO.Counter, glob Perm, glob DPre, glob SRO.Bounder} + /\ DPre.h{1} = ha + /\ invm Perm.m{1} Perm.mi{1} /\ ={c}(SRO.Counter,Cntr)). + + proc; sp; if; auto; sp; if; auto; sp. + exists * x{1}; elim* => m c1 c2 b1 b2. + by call(equiv_sponge_perm c1 m); auto; smt(). + auto; progress. + by rewrite /invm=> x y; rewrite 2!emptyE. + have->//:Pr[Indiff0.Indif(RO, SimSetSize(Simulator, RO), DPre(A)).main() @ &m : res] = + Pr[IdealIndif(BIRO.IRO, Simulator, DRestr(DSetSize(DPre(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline Simulator(FGetSize(RO)).init RO.init Simulator(BIRO.IRO).init + BIRO.IRO.init Gconcl_list.BIRO2.IRO.init; sp. + inline DPre(A, RO, Simulator(FGetSize(RO))).distinguish. + inline DRestr(DSetSize(DPre(A)), BIRO.IRO, Simulator(BIRO.IRO)).distinguish + DSetSize(DPre(A), FC(BIRO.IRO), PC(Simulator(BIRO.IRO))).distinguish; wp; sim. + inline SRO.Bounder(FInit(DFSetSize(FC(BIRO.IRO)))).init + SRO.Bounder(FInit(RO)).init SRO.Counter.init FInit(RO).init + FInit(DFSetSize(FC(BIRO.IRO))).init Cntr.init; sp. + inline SRO.Preimage(A, FInit(RO)).main + SRO.Preimage(A, FInit(DFSetSize(FC(BIRO.IRO)))).main. + inline SRO.Counter.init SRO.Bounder(FInit(RO)).init + SRO.Bounder(FInit(DFSetSize(FC(BIRO.IRO)))).init + FInit(RO).init FInit(DFSetSize(FC(BIRO.IRO))).init ; sp; sim. + seq 1 1 : (={m, glob SRO.Counter, glob DPre, hash} + /\ ={c}(SRO.Counter,Cntr) /\ DPre.h{1} = hash{1} + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); last first. + - if; auto; sp. + exists * m{1}, SRO.Counter.c{1}; elim* => mess c. + by call(equiv_ro_iro c mess); auto; smt(). + conseq(:_==> ={m, glob SRO.Counter, glob SRO.Bounder, glob DPre} + /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); progress. + call(: ={glob SRO.Counter, glob SRO.Bounder, glob DPre} + /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); auto. + + proc; sp; if; auto; sp; if; auto; sp. + exists* x{1}; elim* => a c1 c2 b1 b2. + call(equiv_ro_iro c1 a); auto; smt(). + smt(mem_empty). + have->//=:= SHA3Indiff (DSetSize(DPre(A))) &m _. + move=> F P P_f_ll P_fi_ll F_ll; proc; inline*; auto; sp; auto. + seq 1 : true; auto. + + call (A_ll (SRO.Bounder(FInit(DFSetSize(F)))) _); auto. + by proc; inline*; sp; if; auto; sp; if; auto; sp; call F_ll; auto. + if; auto; sp. + by call F_ll; auto. + qed. + +end section Preimage. + + + +section SecondPreimage. + + declare module A <: SRO.AdvSecondPreimage {-SRO.RO.RO, -SRO.RO.FRO, -SRO.Bounder, -Perm, -Gconcl_list.BIRO2.IRO, -Simulator, -Cntr, -BIRO.IRO, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -D2Pre}. + + declare axiom A_ll (F <: SRO.Oracle {-A}) : islossless F.get => islossless A(F).guess. + + local lemma invm_dom_rng (m mi : (state, state) fmap) : + invm m mi => dom m = rng mi. + proof. + move=>h; rewrite fun_ext=> x; rewrite domE rngE /= eq_iff; have h2 := h x; split. + + move=> m_x_not_None; exists (oget m.[x]); rewrite -h2; move: m_x_not_None. + by case: (m.[x]). + by move=> [] a; rewrite -h2 => ->. + qed. + + local lemma invmC' (m mi : (state, state) fmap) : + invm m mi => invm mi m. + proof. by rewrite /#. qed. + + local lemma invmC (m mi : (state, state) fmap) : + invm m mi <=> invm mi m. + proof. by split;exact invmC'. qed. + + local lemma useful m mi a : + Prefix.invm m mi => ! a \in m => Distr.is_lossless ((bdistr `*` cdistr) \ rng m). + proof. + move=>hinvm nin_dom. + have prod_ll:Distr.is_lossless (bdistr `*` cdistr). + + by rewrite dprod_ll DBlock.dunifin_ll DCapacity.dunifin_ll. + apply dexcepted_ll=>//=;rewrite-prod_ll. + have->:predT = predU (predC (rng m)) (rng m);1:rewrite predCU//=. + rewrite Distr.mu_disjoint 1:predCI//=RField.addrC. + have/=->:=StdOrder.RealOrder.ltr_add2l (mu (bdistr `*` cdistr) (rng m)) 0%r. + rewrite Distr.witness_support/predC. + move:nin_dom;apply absurd=>//=;rewrite negb_exists/==>hyp. + have{hyp}hyp:forall x, rng m x by smt(supp_dprod DBlock.supp_dunifin DCapacity.supp_dunifin). + move:a. + have:=eqEcard (fdom m) (frng m);rewrite leq_card_rng_dom/=. + have->//=:fdom m \subset frng m. + + by move=> x; rewrite mem_fdom mem_frng hyp. + smt(mem_fdom mem_frng). + qed. + + + local equiv equiv_sponge_perm c m : + FInit(CSetSize(Sponge, Perm)).get ~ FInit(DFSetSize(FC(Sponge(Perm)))).get : + ={arg, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c /\ arg{2} = m /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} ==> + ={res, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline FC(Sponge(Perm)).f; sp. + rcondt{2} 1; auto; sp. + call(: ={glob Perm} /\ invm Perm.m{1} Perm.mi{1})=>/=; auto; inline*. + while(={i, n, sa, sc, z, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + + sp; if; auto; sp; if; auto; progress. + rewrite invm_set //=. + by move:H4; rewrite supp_dexcepted. + sp; conseq(:_==> ={i, n, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + while(={xs, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + sp; if; auto; progress. + rewrite invm_set=>//=. + by move:H4; rewrite supp_dexcepted. + qed. + + + clone import Program as Prog2 with + type t <- bool, + op d <- dbool + proof *. + + local equiv equiv_ro_iro c m : + FInit(RO).get ~ FInit(DFSetSize(FC(BIRO.IRO))).get : + ={arg} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + arg{2} = m /\ Cntr.c{2} = c /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} + ==> ={res} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline *; sp; rcondt{2} 1; 1: auto. + swap{2} 1 5; sp; wp 2 1. + conseq(:_==> oget SRO.RO.RO.m{1}.[x{1}] = oget (of_list bs0{2}) /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); 1:by auto. + rcondt{2} 1; 1: auto. + case: (x{1} \in SRO.RO.RO.m{1}). + + rcondf{1} 2; auto. + exists* BIRO.IRO.mp{2}; elim* => mp. + while{2}(bs0{2} = map (fun j => oget BIRO.IRO.mp{2}.[(x0{2},j)]) (range 0 i{2}) + /\ n0{2} = size_out /\ x0{2} \in SRO.RO.RO.m{1} /\ 0 <= i{2} <= size_out + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ BIRO.IRO.mp{2} = mp) + (size_out - i{2}); auto. + - sp; rcondf 1; auto; 1: smt(). + progress. + * have/=<-:= map_rcons (fun (j : int) => oget BIRO.IRO.mp{hr}.[(x0{hr}, j)]) (range 0 i{hr}) i{hr}. + by rewrite !rangeSr //=. + * smt(). + * smt(). + * smt(). + progress. + - by rewrite range_geq. + - smt(size_out_gt0). + - smt(). + - exact(dout_ll). + - have[] h[#] h1 h2 := H. + have->:i_R = size_out by smt(). + have<-:=h2 _ H3. + smt(to_listK). + rcondt{1} 2; 1: auto; wp =>/=. + exists* BIRO.IRO.mp{2}; elim* => mp. + conseq(:_==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress=>//=. + + by rewrite get_set_sameE /=; smt(to_listK take_oversize spec_dout). + + move:H8; rewrite mem_set=>[][]//=h; 1:rewrite H3=>//=. + - by have []h1 []h2 h3:= H2; have->//:=h1 _ h. + by move:h => <<-; rewrite H6 //=. + + rewrite mem_set//=; have[]//=h:= H5 _ _ H11; left. + have []h1 []->//=:= H2. + by exists i0=>//=. + + move:H7; rewrite take_oversize 1:spec_dout//= => H7. + move:H10; rewrite mem_set. + case(m \in SRO.RO.RO.m{1})=>//=h. + - rewrite get_set_neqE; 1:smt(). + have []h1 []h2 ->//=:= H2. + by apply eq_in_map=> j;rewrite mem_range=>[][]hj1 hj2/=; rewrite H4//=h1//=. + by move=><<-; rewrite get_set_eqE//=. + alias{1} 1 l = [<:bool>]. + transitivity {1} { + l <@ Sample.sample(size_out); + r <- oget (of_list l); + } + (={glob SRO.RO.RO, x} ==> ={glob SRO.RO.RO, r}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x{1} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + inline*; sp; wp. + rnd to_list (fun x => oget (of_list x)); auto; progress. + - smt(spec_dout supp_dlist to_listK spec2_dout size_out_gt0). + - rewrite -dout_equal_dlist dmap1E; apply mu_eq=> x/=. + smt(to_listK). + - rewrite-dout_equal_dlist supp_dmap; smt(dout_full). + smt(to_listK). + wp=>/=. + conseq(:_==> i{2} = size_out /\ size l{1} = size_out /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => (l0, j) \in BIRO.IRO.mp{2}) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => BIRO.IRO.mp{2}.[(l0, j)] = mp.[(l0, j)]) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in BIRO.IRO.mp{2} => ((l0, j) \in mp) \/ (l0 = x0{2} /\ 0 <= j < i{2})) /\ + (forall (j : int), 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + transitivity{1} { + l <@ LoopSnoc.sample(size_out); + } + (={glob SRO.RO.RO} ==> ={glob SRO.RO.RO, l}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x0{2} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + i{2} = size_out /\ size l{1} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + by call Sample_LoopSnoc_eq; auto. + inline*; sp; wp. + conseq(:_==> i{2} = size_out /\ size l0{1} = i{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress. + + smt(take_oversize). + + smt(take_oversize). + while(0 <= i{2} <= size_out /\ size l0{1} = i{2} /\ n0{2} = size_out /\ + ={i} /\ n{1} = n0{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})). + + sp; wp=> //=. + rcondt{2} 1; 1:auto; progress. + - have[]h1 [] h2 h3 := H1. + have:=h2 x0{hr}; rewrite H2/= negb_exists/= =>/(_ (size bs0{hr})). + rewrite size_ge0 H9/=; apply absurd =>/= h. + by have //=:= H5 _ _ h. + wp; rnd; auto; progress. + - smt(size_ge0). + - smt(). + - by rewrite size_cat/=. + - by rewrite mem_set; left; rewrite H3. + - rewrite get_setE (H4 _ _ H11). + have/#: !(l1, j) = (x0{2}, size bs0{2}). + move:H2; apply absurd=> //=[#] <<- ->>. + have[] h1 [] h2 h3 := H1. + by apply h2; smt(). + - move:H11; rewrite mem_set. + case((l1, j) \in BIRO.IRO.mp{2})=>//= h; 1: smt(). + by move=> [#] <<- ->> //=; rewrite size_ge0; smt(). + - rewrite mem_set. + case(j = size bs0{2})=>//=. + move=> h; rewrite h /=; have {H12} H12 {h} : j < size bs0{2} by smt(). + by apply H6. + - by rewrite cats1 get_set_sameE oget_some. + - rewrite get_set_sameE oget_some H7 rangeSr. + rewrite !size_map 1:size_ge0. + rewrite (size_map _ (range 0 (size bs0{2}))) size_range /=. + rewrite ler_maxr 1:size_ge0 map_rcons /=get_set_sameE oget_some; congr. + apply eq_in_map=> j. + rewrite mem_range /==> [] [] hj1 hj2. + by rewrite get_set_neqE //=; smt(). + auto; progress. + + smt(size_out_gt0). + + smt(). + + smt(). + + by rewrite range_geq. + smt(). + qed. + + lemma Sponge_second_preimage_resistant &m mess : + (D2Pre.m2{m} = mess) => + Pr[SRO.SecondPreimage(A, FM(CSetSize(Sponge), Perm)).main(mess) @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma + 1)%r / (2%r ^ size_out). + proof. + move=> init_mess. + rewrite -(doutE1 witness). + rewrite(second_preimage_resistant_if_indifferentiable A A_ll (CSetSize(Sponge)) Perm &m mess init_mess). + exists (SimSetSize(Simulator)); split. + + by move=> F _; proc; inline*; auto. + have->:Pr[Indiff0.Indif(CSetSize(Sponge, Perm), Perm, D2Pre(A)).main() @ &m : res] = + Pr[RealIndif(Sponge, Perm, DRestr(DSetSize(D2Pre(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline Perm.init CSetSize(Sponge, Perm).init Sponge(Perm).init + FC(Sponge(Perm)).init; sp. + inline D2Pre(A, CSetSize(Sponge, Perm), Perm).distinguish. + inline DRestr(DSetSize(D2Pre(A)), Sponge(Perm), Perm).distinguish + DSetSize(D2Pre(A), FC(Sponge(Perm)), PC(Perm)).distinguish Cntr.init. + inline SRO.SecondPreimage(A, FInit(CSetSize(Sponge, Perm))).main + SRO.SecondPreimage(A, FInit(DFSetSize(FC(Sponge(Perm))))).main. + inline SRO.Bounder(FInit(CSetSize(Sponge, Perm))).init + SRO.Bounder(FInit(DFSetSize(FC(Sponge(Perm))))).init + SRO.Counter.init FInit(DFSetSize(FC(Sponge(Perm)))).init + FInit(CSetSize(Sponge, Perm)).init. + wp; sp; sim. + seq 1 1 : (={m1, m2, glob SRO.Counter, glob Perm} + /\ invm Perm.m{1} Perm.mi{1} + /\ ={c}(SRO.Counter,Cntr)); last first. + - if; auto; sp. + case(SRO.Counter.c{1} + ((size m2{1} + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 < limit); last first. + * rcondf{1} 2; 1: by auto; inline*; auto; conseq(: _ ==> true); auto. + rcondf{2} 2; 1: by auto; inline*; auto; conseq(: _ ==> true); auto. + auto; inline*; auto; sp; conseq(: _ ==> true); auto. + if{2}; sp; auto; sim. + while{1}(invm Perm.m{1} Perm.mi{1}) (((size_out + r - 1) %/ r)-i{1}). + + auto; sp; if; auto. + - sp; if ;auto; progress. + * exact (useful _ _ _ H H2). + * rewrite invm_set=>//=. + by move:H4; rewrite supp_dexcepted. + * smt(). + smt(). + smt(). + conseq(:_==> invm Perm.m{1} Perm.mi{1}); 1:smt(). + while{1}(invm Perm.m{1} Perm.mi{1})(size xs{1}). + + move=> _ z; auto; sp; if; auto; progress. + * exact (useful _ _ _ H H1). + * rewrite invm_set=>//=. + by move:H3; rewrite supp_dexcepted. + * smt(). + smt(). + auto; smt(size_ge0 size_eq0). + rcondt{1} 2; first by auto; inline*; auto; conseq(:_==> true); auto. + rcondt{2} 2; first by auto; inline*; auto; conseq(:_==> true); auto. + sim. + exists* m1{1}, m2{1}; elim* => a1 a2 c1 c2. + call (equiv_sponge_perm (c2 + ((size a1 + 1) %/ r + 1) + max ((size_out + r - 1) %/ r - 1) 0) a2). + auto; call (equiv_sponge_perm c2 a1); auto; progress. + smt(List.size_ge0 divz_ge0 gt0_r). + smt(List.size_ge0 divz_ge0 gt0_r). + call(: ={glob SRO.Counter, glob Perm, glob SRO.Bounder} + /\ invm Perm.m{1} Perm.mi{1} /\ ={c}(SRO.Counter,Cntr)). + + proc; sp; if; auto; sp; if; auto; sp. + exists * x{1}; elim* => m c1 c2 b1 b2. + by call(equiv_sponge_perm c1 m); auto; smt(). + inline*; auto; progress. + by rewrite /invm=> x y; rewrite 2!emptyE. + have->:Pr[Indiff0.Indif(RO, SimSetSize(Simulator, RO), D2Pre(A)).main() @ &m : res] = + Pr[IdealIndif(BIRO.IRO, Simulator, DRestr(DSetSize(D2Pre(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline Simulator(FGetSize(RO)).init RO.init Simulator(BIRO.IRO).init + BIRO.IRO.init Gconcl_list.BIRO2.IRO.init; sp. + inline D2Pre(A, RO, Simulator(FGetSize(RO))).distinguish. + inline DRestr(DSetSize(D2Pre(A)), BIRO.IRO, Simulator(BIRO.IRO)).distinguish + DSetSize(D2Pre(A), FC(BIRO.IRO), PC(Simulator(BIRO.IRO))).distinguish; wp; sim. + inline SRO.Bounder(FInit(DFSetSize(FC(BIRO.IRO)))).init + SRO.Bounder(FInit(RO)).init SRO.Counter.init FInit(RO).init + FInit(DFSetSize(FC(BIRO.IRO))).init Cntr.init; sp. + inline SRO.SecondPreimage(A, FInit(RO)).main + SRO.SecondPreimage(A, FInit(DFSetSize(FC(BIRO.IRO)))).main. + inline SRO.Bounder(FInit(RO)).init + SRO.Bounder(FInit(DFSetSize(FC(BIRO.IRO)))).init SRO.Counter.init + FInit(RO).init FInit(DFSetSize(FC(BIRO.IRO))).init. + sp; sim. + seq 1 1 : (={m1, m2, glob SRO.Counter} + /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); last first. + - if; auto; sp. + case: (SRO.Counter.c{1} + ((size m2{1} + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 < limit); last first. + * rcondf{1} 2; first by auto; inline*; auto. + rcondf{2} 2; first auto; inline*; auto; sp. + + rcondt 1; first by auto; smt(). + by sp; rcondt 1; auto; conseq(:_==> true); auto. + inline*;sp; auto. + rcondt{2} 1; first by auto; smt(). + conseq(:_==> true); first smt(dout_ll). + sp; rcondt{2} 1; auto; conseq(:_==> true); auto. + by while{2}(true)(n0{2}-i{2}); auto; 1:(sp; if; auto); smt(dbool_ll). + rcondt{1} 2; first by auto; inline*; auto. + rcondt{2} 2; first auto; inline*; auto; sp. + + rcondt 1; first by auto; smt(). + by sp; rcondt 1; auto; conseq(:_==> true); auto. + sim. + exists* m1{1}, m2{1}; elim*=> a1 a2 c1 c2. + call(equiv_ro_iro (c2 + ((size a1 + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0) a2). + auto; call(equiv_ro_iro c2 a1); auto; smt(). + call(: ={glob SRO.Counter, glob SRO.Bounder} /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); auto. + + proc; sp; if; auto; sp; if; auto; sp. + exists* x{1}; elim* => a c1 c2 b1 b2. + call(equiv_ro_iro c1 a); auto; smt(). + smt(mem_empty). + have->//=:= SHA3Indiff (DSetSize(D2Pre(A))) &m _. + move=> F P P_f_ll P_fi_ll F_ll; proc; inline*; auto; sp. + seq 1 : true; auto. + + call (A_ll (SRO.Bounder(FInit(DFSetSize(F)))) _); auto. + by proc; inline*; sp; if; auto; sp; if; auto; sp; call F_ll; auto. + if; auto; sp. + seq 1 : true; auto. + + by call F_ll; auto. + sp; if; auto; sp; call F_ll; auto. + qed. + +end section SecondPreimage. + + + +section Collision. + + declare module A <: SRO.AdvCollision {-SRO.RO.RO, -SRO.RO.FRO, -SRO.Bounder, -Perm, -Gconcl_list.BIRO2.IRO, -Simulator, -Cntr, -BIRO.IRO, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator}. + + declare axiom A_ll (F <: SRO.Oracle {-A}) : islossless F.get => islossless A(F).guess. + + local lemma invm_dom_rng (m mi : (state, state) fmap) : + invm m mi => dom m = rng mi. + proof. + move=>h; rewrite fun_ext=> x; rewrite domE rngE /= eq_iff; have h2 := h x; split. + + move=> m_x_not_None; exists (oget m.[x]); rewrite -h2; move: m_x_not_None. + by case: (m.[x]). + by move=> [] a; rewrite -h2 => ->. + qed. + + local lemma invmC' (m mi : (state, state) fmap) : + invm m mi => invm mi m. + proof. by rewrite /#. qed. + + local lemma invmC (m mi : (state, state) fmap) : + invm m mi <=> invm mi m. + proof. by split;exact invmC'. qed. + + local lemma useful m mi a : + Prefix.invm m mi => ! a \in m => Distr.is_lossless ((bdistr `*` cdistr) \ rng m). + proof. + move=>hinvm nin_dom. + have prod_ll:Distr.is_lossless (bdistr `*` cdistr). + + by rewrite dprod_ll DBlock.dunifin_ll DCapacity.dunifin_ll. + apply dexcepted_ll=>//=;rewrite-prod_ll. + have->:predT = predU (predC (rng m)) (rng m);1:rewrite predCU//=. + rewrite Distr.mu_disjoint 1:predCI//=RField.addrC. + have/=->:=StdOrder.RealOrder.ltr_add2l (mu (bdistr `*` cdistr) (rng m)) 0%r. + rewrite Distr.witness_support/predC. + move:nin_dom;apply absurd=>//=;rewrite negb_exists/==>hyp. + have{hyp}hyp:forall x, rng m x by smt(supp_dprod DBlock.supp_dunifin DCapacity.supp_dunifin). + move:a. + have:=eqEcard (fdom m) (frng m);rewrite leq_card_rng_dom/=. + have->//=:fdom m \subset frng m. + + by move=> x; rewrite mem_fdom mem_frng hyp. + smt(mem_fdom mem_frng). + qed. + + + local equiv equiv_sponge_perm c m : + FInit(CSetSize(Sponge, Perm)).get ~ FInit(DFSetSize(FC(Sponge(Perm)))).get : + ={arg, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c /\ arg{2} = m /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} ==> + ={res, glob Perm} /\ invm Perm.m{1} Perm.mi{1} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline FC(Sponge(Perm)).f; sp. + rcondt{2} 1; auto; sp. + call(: ={glob Perm} /\ invm Perm.m{1} Perm.mi{1})=>/=; auto; inline*. + while(={i, n, sa, sc, z, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + + sp; if; auto; sp; if; auto; progress. + rewrite invm_set //=. + by move:H4; rewrite supp_dexcepted. + sp; conseq(:_==> ={i, n, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + while(={xs, sa, sc, glob Perm} /\ invm Perm.m{1} Perm.mi{1}); auto. + sp; if; auto; progress. + rewrite invm_set=>//=. + by move:H4; rewrite supp_dexcepted. + qed. + + clone import Program as Prog3 with + type t <- bool, + op d <- dbool + proof *. + + local equiv equiv_ro_iro c m : + FInit(RO).get ~ FInit(DFSetSize(FC(BIRO.IRO))).get : + ={arg} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + arg{2} = m /\ Cntr.c{2} = c /\ + (Cntr.c + ((size arg + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0 <= limit){2} + ==> ={res} /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ + Cntr.c{2} = c + ((size m + 1) %/ Common.r + 1) + + max ((size_out + Common.r - 1) %/ Common.r - 1) 0. + proof. + proc; inline *; sp; rcondt{2} 1; 1: auto. + swap{2} 1 5; sp; wp 2 1. + conseq(:_==> oget SRO.RO.RO.m{1}.[x{1}] = oget (of_list bs0{2}) /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); 1:by auto. + rcondt{2} 1; 1: auto. + case: (x{1} \in SRO.RO.RO.m{1}). + + rcondf{1} 2; auto. + exists* BIRO.IRO.mp{2}; elim* => mp. + while{2}(bs0{2} = map (fun j => oget BIRO.IRO.mp{2}.[(x0{2},j)]) (range 0 i{2}) + /\ n0{2} = size_out /\ x0{2} \in SRO.RO.RO.m{1} /\ 0 <= i{2} <= size_out + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ BIRO.IRO.mp{2} = mp) + (size_out - i{2}); auto. + - sp; rcondf 1; auto; 1: smt(). + progress. + * have/=<-:= map_rcons (fun (j : int) => oget BIRO.IRO.mp{hr}.[(x0{hr}, j)]) (range 0 i{hr}) i{hr}. + by rewrite !rangeSr //=. + * smt(). + * smt(). + * smt(). + progress. + - by rewrite range_geq. + - smt(size_out_gt0). + - smt(). + - exact(dout_ll). + - have[] h[#] h1 h2 := H. + have->:i_R = size_out by smt(). + have<-:=h2 _ H3. + smt(to_listK). + rcondt{1} 2; 1: auto; wp =>/=. + exists* BIRO.IRO.mp{2}; elim* => mp. + conseq(:_==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress=>//=. + + by rewrite get_set_sameE /=; smt(to_listK take_oversize spec_dout). + + move:H8; rewrite mem_set=>[][]//=h; 1:rewrite H3=>//=. + - by have []h1 []h2 h3:= H2; have->//:=h1 _ h. + by move:h => <<-; rewrite H6 //=. + + rewrite mem_set//=; have[]//=h:= H5 _ _ H11; left. + have []h1 []->//=:= H2. + by exists i0=>//=. + + move:H7; rewrite take_oversize 1:spec_dout//= => H7. + move:H10; rewrite mem_set. + case(m \in SRO.RO.RO.m{1})=>//=h. + - rewrite get_set_neqE; 1:smt(). + have []h1 []h2 ->//=:= H2. + by apply eq_in_map=> j;rewrite mem_range=>[][]hj1 hj2/=; rewrite H4//=h1//=. + by move=><<-; rewrite get_set_eqE//=. + alias{1} 1 l = [<:bool>]. + transitivity {1} { + l <@ Sample.sample(size_out); + r <- oget (of_list l); + } + (={glob SRO.RO.RO, x} ==> ={glob SRO.RO.RO, r}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x{1} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + same_ro SRO.RO.RO.m{1} mp /\ i{2} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} (to_list r{1}) = bs0{2} /\ + take i{2} (to_list r{1}) = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + inline*; sp; wp. + rnd to_list (fun x => oget (of_list x)); auto; progress. + - smt(spec_dout supp_dlist to_listK spec2_dout size_out_gt0). + - rewrite -dout_equal_dlist dmap1E; apply mu_eq=> x/=. + smt(to_listK). + - rewrite-dout_equal_dlist supp_dmap; smt(dout_full). + smt(to_listK). + wp=>/=. + conseq(:_==> i{2} = size_out /\ size l{1} = size_out /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => (l0, j) \in BIRO.IRO.mp{2}) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in mp => BIRO.IRO.mp{2}.[(l0, j)] = mp.[(l0, j)]) /\ + (forall (l0 : bool list) (j : int), + (l0, j) \in BIRO.IRO.mp{2} => ((l0, j) \in mp) \/ (l0 = x0{2} /\ 0 <= j < i{2})) /\ + (forall (j : int), 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + + have[]//=h h1:=to_listK (oget (of_list l_L)) l_L; rewrite h1//==> {h1 h}. + smt(spec2_dout). + transitivity{1} { + l <@ LoopSnoc.sample(size_out); + } + (={glob SRO.RO.RO} ==> ={glob SRO.RO.RO, l}) + (x{1} = x0{2} /\ i{2} = 0 /\ n0{2} = size_out /\ mp = BIRO.IRO.mp{2} /\ + same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2} /\ x0{2} \notin SRO.RO.RO.m{1} /\ + bs0{2} = [] + ==> + i{2} = size_out /\ size l{1} = size_out /\ + (forall (l,j), (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall (l,j), (l,j) \in mp => BIRO.IRO.mp{2}.[(l,j)] = mp.[(l,j)]) /\ + (forall (l,j), (l,j) \in BIRO.IRO.mp{2} => (l,j) \in mp \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2},j) \in BIRO.IRO.mp{2}) /\ + take i{2} l{1} = bs0{2} /\ + take i{2} l{1} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); + progress. + + smt(). + + by call Sample_LoopSnoc_eq; auto. + inline*; sp; wp. + conseq(:_==> i{2} = size_out /\ size l0{1} = i{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})); progress. + + smt(take_oversize). + + smt(take_oversize). + while(0 <= i{2} <= size_out /\ size l0{1} = i{2} /\ n0{2} = size_out /\ + ={i} /\ n{1} = n0{2} /\ + same_ro SRO.RO.RO.m{1} mp /\ x0{2} \notin SRO.RO.RO.m{1} /\ + (forall l j, (l,j) \in mp => (l,j) \in BIRO.IRO.mp{2}) /\ + (forall l j, (l,j) \in mp => BIRO.IRO.mp{2}.[(l, j)] = mp.[(l, j)]) /\ + (forall l j, (l, j) \in BIRO.IRO.mp{2} => ((l, j) \in mp) \/ (l = x0{2} /\ 0 <= j < i{2})) /\ + (forall j, 0 <= j < i{2} => (x0{2}, j) \in BIRO.IRO.mp{2}) /\ + l0{1} = bs0{2} /\ bs0{2} = + map (fun (j : int) => oget BIRO.IRO.mp{2}.[(x0{2}, j)]) (range 0 i{2})). + + sp; wp=> //=. + rcondt{2} 1; 1:auto; progress. + - have[]h1 [] h2 h3 := H1. + have:=h2 x0{hr}; rewrite H2/= negb_exists/= =>/(_ (size bs0{hr})). + rewrite size_ge0 H9/=; apply absurd =>/= h. + by have //=:= H5 _ _ h. + wp; rnd; auto; progress. + - smt(size_ge0). + - smt(). + - by rewrite size_cat/=. + - by rewrite mem_set; left; rewrite H3. + - rewrite get_setE (H4 _ _ H11). + have/#: !(l1, j) = (x0{2}, size bs0{2}). + move:H2; apply absurd=> //=[#] <<- ->>. + have[] h1 [] h2 h3 := H1. + by apply h2; smt(). + - move:H11; rewrite mem_set. + case((l1, j) \in BIRO.IRO.mp{2})=>//= h; 1: smt(). + by move=> [#] <<- ->> //=; rewrite size_ge0; smt(). + - rewrite mem_set. + case(j = size bs0{2})=>//=. + move=> h; rewrite h /=; have {H12} H12 {h} : j < size bs0{2} by smt(). + by apply H6. + - by rewrite cats1 get_set_sameE oget_some. + - rewrite get_set_sameE oget_some H7 rangeSr. + rewrite !size_map 1:size_ge0. + rewrite (size_map _ (range 0 (size bs0{2}))) size_range /=. + rewrite ler_maxr 1:size_ge0 map_rcons /=get_set_sameE oget_some; congr. + apply eq_in_map=> j. + rewrite mem_range /==> [] [] hj1 hj2. + by rewrite get_set_neqE //=; smt(). + auto; progress. + + smt(size_out_gt0). + + smt(). + + smt(). + + by rewrite range_geq. + smt(). + qed. + + lemma Sponge_coll_resistant &m : + Pr[SRO.Collision(A, FM(CSetSize(Sponge), Perm)).main() @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma * (sigma - 1) + 2)%r / 2%r / (2%r ^ size_out). + proof. + rewrite -(doutE1 witness). + rewrite (coll_resistant_if_indifferentiable A A_ll (CSetSize(Sponge)) Perm &m). + exists (SimSetSize(Simulator)); split. + + by move=> F _; proc; inline*; auto. + have->:Pr[Indiff0.Indif(CSetSize(Sponge, Perm), Perm, DColl(A)).main() @ &m : res] = + Pr[RealIndif(Sponge, Perm, DRestr(DSetSize(DColl(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline Perm.init CSetSize(Sponge, Perm).init Sponge(Perm).init + FC(Sponge(Perm)).init; sp. + inline DColl(A, CSetSize(Sponge, Perm), Perm).distinguish. + inline DRestr(DSetSize(DColl(A)), Sponge(Perm), Perm).distinguish + DSetSize(DColl(A), FC(Sponge(Perm)), PC(Perm)).distinguish Cntr.init; wp; sp; sim. + seq 2 2 : (={m1, m2, glob SRO.Counter, glob Perm} + /\ invm Perm.m{1} Perm.mi{1} + /\ ={c}(SRO.Counter,Cntr)); last first. + - if; auto; sp. + case(SRO.Counter.c{1} + ((size m2{1} + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 < limit); last first. + * rcondf{1} 2; 1: by auto; inline*; auto; conseq(: _ ==> true); auto. + rcondf{2} 2; 1: by auto; inline*; auto; conseq(: _ ==> true); auto. + auto; inline*; auto; sp; conseq(: _ ==> true); auto. + if{2}; sp; auto; sim. + while{1}(invm Perm.m{1} Perm.mi{1}) (((size_out + r - 1) %/ r)-i{1}). + + auto; sp; if; auto. + - sp; if ;auto; progress. + * exact (useful _ _ _ H H2). + * rewrite invm_set=>//=. + by move:H4; rewrite supp_dexcepted. + * smt(). + smt(). + smt(). + conseq(:_==> invm Perm.m{1} Perm.mi{1}); 1:smt(). + while{1}(invm Perm.m{1} Perm.mi{1})(size xs{1}). + + move=> _ z; auto; sp; if; auto; progress. + * exact (useful _ _ _ H H1). + * rewrite invm_set=>//=. + by move:H3; rewrite supp_dexcepted. + * smt(). + smt(). + auto; smt(size_ge0 size_eq0). + rcondt{1} 2; first by auto; inline*; auto; conseq(:_==> true); auto. + rcondt{2} 2; first by auto; inline*; auto; conseq(:_==> true); auto. + sim. + exists* m1{1}, m2{1}; elim* => a1 a2 c1 c2. + call (equiv_sponge_perm (c2 + ((size a1 + 1) %/ r + 1) + max ((size_out + r - 1) %/ r - 1) 0) a2). + auto; call (equiv_sponge_perm c2 a1); auto; progress. + smt(List.size_ge0 divz_ge0 gt0_r). + smt(List.size_ge0 divz_ge0 gt0_r). + call(: ={glob SRO.Counter, glob Perm, glob SRO.Bounder} + /\ invm Perm.m{1} Perm.mi{1} /\ ={c}(SRO.Counter,Cntr)). + + proc; sp; if; auto; sp; if; auto; sp. + exists * x{1}; elim* => m c1 c2 b1 b2. + by call(equiv_sponge_perm c1 m); auto; smt(). + inline*; auto; progress. + by rewrite /invm=> x y; rewrite 2!emptyE. + have->:Pr[Indiff0.Indif(RO, SimSetSize(Simulator, RO), DColl(A)).main() @ &m : res] = + Pr[IdealIndif(BIRO.IRO, Simulator, DRestr(DSetSize(DColl(A)))).main() @ &m : res]. + + byequiv=>//=; proc. + inline Simulator(FGetSize(RO)).init RO.init Simulator(BIRO.IRO).init + BIRO.IRO.init Gconcl_list.BIRO2.IRO.init; sp. + inline DColl(A, RO, Simulator(FGetSize(RO))).distinguish. + inline DRestr(DSetSize(DColl(A)), BIRO.IRO, Simulator(BIRO.IRO)).distinguish + DSetSize(DColl(A), FC(BIRO.IRO), PC(Simulator(BIRO.IRO))).distinguish; wp; sim. + inline SRO.Bounder(FInit(DFSetSize(FC(BIRO.IRO)))).init + SRO.Bounder(FInit(RO)).init SRO.Counter.init FInit(RO).init + FInit(DFSetSize(FC(BIRO.IRO))).init Cntr.init; sp. + seq 1 1 : (={m1, m2, glob SRO.Counter} + /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); last first. + - if; auto; sp. + case: (SRO.Counter.c{1} + ((size m2{1} + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0 < limit); last first. + * rcondf{1} 2; first by auto; inline*; auto. + rcondf{2} 2; first auto; inline*; auto; sp. + + rcondt 1; first by auto; smt(). + by sp; rcondt 1; auto; conseq(:_==> true); auto. + inline*;sp; auto. + rcondt{2} 1; first by auto; smt(). + conseq(:_==> true); first smt(dout_ll). + sp; rcondt{2} 1; auto; conseq(:_==> true); auto. + by while{2}(true)(n0{2}-i{2}); auto; 1:(sp; if; auto); smt(dbool_ll). + rcondt{1} 2; first by auto; inline*; auto. + rcondt{2} 2; first auto; inline*; auto; sp. + + rcondt 1; first by auto; smt(). + by sp; rcondt 1; auto; conseq(:_==> true); auto. + sim. + exists* m1{1}, m2{1}; elim*=> a1 a2 c1 c2. + call(equiv_ro_iro (c2 + ((size a1 + 1) %/ r + 1) + + max ((size_out + r - 1) %/ r - 1) 0) a2). + auto; call(equiv_ro_iro c2 a1); auto; smt(). + call(: ={glob SRO.Counter, glob SRO.Bounder} /\ ={c}(SRO.Counter,Cntr) + /\ same_ro SRO.RO.RO.m{1} BIRO.IRO.mp{2}); auto. + + proc; sp; if; auto; sp; if; auto; sp. + exists* x{1}; elim* => a c1 c2 b1 b2. + call(equiv_ro_iro c1 a); auto; smt(). + smt(mem_empty). + have->//=:= SHA3Indiff (DSetSize(DColl(A))) &m _. + move=> F P P_f_ll P_fi_ll F_ll; proc; inline*; auto; sp. + seq 1 : true; auto. + + call (A_ll (SRO.Bounder(FInit(DFSetSize(F)))) _); auto. + by proc; inline*; sp; if; auto; sp; if; auto; sp; call F_ll; auto. + if; auto; sp. + seq 1 : true; auto. + + by call F_ll; auto. + sp; if; auto; sp; call F_ll; auto. + qed. + +end section Collision. + +module X (F : SRO.Oracle) = { + proc get (bl : bool list) = { + var r; + r <@ F.get(bl ++ [ false ; true ]); + return r; + } +}. + +module AdvCollisionSHA3 (A : SRO.AdvCollision) (F : SRO.Oracle) = { + proc guess () = { + var m1, m2; + (m1, m2) <@ A(X(F)).guess(); + return (m1 ++ [ false ; true ], m2 ++ [ false ; true ]); + } +}. + +section SHA3_Collision. + + declare module A <: SRO.AdvCollision {-SRO.RO.RO, -SRO.RO.FRO, -SRO.Bounder, -Perm, -Gconcl_list.BIRO2.IRO, -Simulator, -Cntr, -BIRO.IRO, -F.RO, -F.FRO, -Redo, -C, -Gconcl.S, -BlockSponge.BIRO.IRO, -BlockSponge.C, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator}. + + declare axiom A_ll (F <: SRO.Oracle {-A}) : islossless F.get => islossless A(F).guess. + + lemma SHA3_coll_resistant &m : + Pr[SRO.Collision(AdvCollisionSHA3(A), FM(CSetSize(Sponge), Perm)).main() @ &m : res] <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + + (4 * limit ^ 2)%r / (2 ^ c)%r + + (sigma * (sigma - 1) + 2)%r / 2%r / (2%r ^ size_out). + proof. + apply (Sponge_coll_resistant (AdvCollisionSHA3(A)) _ &m). + by move=> F F_ll; proc; inline*; call(A_ll (X(F)) _); auto; proc; call F_ll; auto. + qed. + + +end section SHA3_Collision. diff --git a/sha3/proof/SHA3_OIndiff.ec b/sha3/proof/SHA3_OIndiff.ec new file mode 100644 index 0000000..34f6679 --- /dev/null +++ b/sha3/proof/SHA3_OIndiff.ec @@ -0,0 +1,259 @@ +require import AllCore List Int IntDiv StdOrder Distr SmtMap FSet. + +require import Common Sponge. import BIRO. +require (*--*) SLCommon Gconcl_list BlockSponge. +require import SHA3Indiff. + +(* FIX: would be nicer to define limit at top-level and then clone + BlockSponge with it - so BlockSponge would then clone lower-level + theories with it + +op limit : {int | 0 < limit} as gt0_max_limit. +*) + +require (****) OptionIndifferentiability. + +clone import OptionIndifferentiability as OIndif with + type p <- state, + type f_out <- bool list, + type f_in <- bool list * int +proof *. + + +module FSome (F : FUNCTIONALITY) : OFUNCTIONALITY = { + proc init = F.init + proc f (x: bool list * int) : bool list option = { + var z; + z <@ F.f(x); + return Some z; + } +}. + +module PSome (P : PRIMITIVE) : OPRIMITIVE = { + proc init = P.init + proc f (x : state) : state option = { + var z; + z <@ P.f(x); + return Some z; + } + proc fi (x: state) : state option = { + var z; + z <@ P.fi(x); + return Some z; + } +}. + +module Poget (P : ODPRIMITIVE) : DPRIMITIVE = { + proc f (x : state) : state = { + var z; + z <@ P.f(x); + return oget z; + } + proc fi (x: state) : state = { + var z; + z <@ P.fi(x); + return oget z; + } +}. + +module (CSome (C : CONSTRUCTION) : OCONSTRUCTION) (P : ODPRIMITIVE) = FSome(C(Poget(P))). + +module OSimulator (F : ODFUNCTIONALITY) = { + proc init() = { + Simulator.m <- empty; + Simulator.mi <- empty; + Simulator.paths <- empty.[c0 <- ([],b0)]; + Gconcl_list.BIRO2.IRO.init(); + } + proc f (x : state) : state option = { + var p,v,z,q,k,cs,y,y1,y2,o; + if (x \notin Simulator.m) { + if (x.`2 \in Simulator.paths) { + (p,v) <- oget Simulator.paths.[x.`2]; + z <- []; + (q,k) <- parse (rcons p (v +^ x.`1)); + if (valid q) { + o <@ F.f(oget (unpad_blocks q), k * r); + cs <- oget o; + z <- bits2blocks cs; + } else { + z <@ Gconcl_list.BIRO2.IRO.f(q,k); + } + y1 <- last b0 z; + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + Simulator.m.[x] <- y; + Simulator.mi.[y] <- x; + if (x.`2 \in Simulator.paths) { + (p,v) <-oget Simulator.paths.[x.`2]; + Simulator.paths.[y2] <- (rcons p (v +^ x.`1),y.`1); + } + } else { + y <- oget Simulator.m.[x]; + } + return Some y; + } + proc fi (x : state) : state option = { + var y,y1,y2; + if (! x \in Simulator.mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + Simulator.mi.[x] <- y; + Simulator.m.[y] <- x; + } else { + y <- oget Simulator.mi.[x]; + } + return Some y; + } +}. + + +module Counter = { + var c : int + proc init () = { + c <- 0; + } +}. + +op increase_counter c (l : 'a list) n = + c + ((size l + 1) %/ r + 1) + max ((n + r - 1) %/ r - 1) 0. + + +module OFC (F : ODFUNCTIONALITY) = { + proc init () = { + Counter.init(); + } + proc f (l : bool list, k : int) : bool list option = { + var o <- None; + if (increase_counter Counter.c l k <= limit) { + o <@ F.f(l,k); + Counter.c <- increase_counter Counter.c l k; + } + return o; + } +}. + +module OPC (P : ODPRIMITIVE) = { + proc init () = {} + proc f (x : state) : state option = { + var o <- None; + if (Counter.c + 1 <= limit) { + o <@ P.f(x); + Counter.c <- Counter.c + 1; + } + return o; + } + proc fi (x : state) : state option = { + var o <- None; + if (Counter.c + 1 <= limit) { + o <@ P.fi(x); + Counter.c <- Counter.c + 1; + } + return o; + } +}. + + +module ODRestr (D : ODISTINGUISHER) (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + proc distinguish () = { + var b; + Counter.init(); + b <@ D(OFC(F),OPC(P)).distinguish(); + return b; + } +}. + +section. +declare module Dist <: + ODISTINGUISHER {-Perm, -Gconcl_list.SimLast, -IRO, -Cntr, -BlockSponge.BIRO.IRO, -Simulator, -BlockSponge.C, -Gconcl.S, -SLCommon.F.RO, -SLCommon.F.FRO, -SLCommon.Redo, -SLCommon.C, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator}. + + +local module DFSome (F : DFUNCTIONALITY) : ODFUNCTIONALITY = { + proc f (x: bool list * int) : bool list option = { + var z; + z <@ F.f(x); + return Some z; + } +}. + +module DPSome (P : DPRIMITIVE) : ODPRIMITIVE = { + proc f (x : state) : state option = { + var z; + z <@ P.f(x); + return Some z; + } + proc fi (x: state) : state option = { + var z; + z <@ P.fi(x); + return Some z; + } +}. + +local module (OD (D : ODISTINGUISHER) : DISTINGUISHER) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish () = { + var b; + Counter.init(); + b <@ D(OFC(DFSome(F)),OPC(DPSome(P))).distinguish(); + return b; + } +}. + +lemma SHA3OIndiff + (Dist <: ODISTINGUISHER{-Counter, -Perm, -IRO, -BlockSponge.BIRO.IRO, -Cntr, -Simulator, -Gconcl_list.SimLast(Gconcl.S), -BlockSponge.C, -Gconcl.S, -SLCommon.F.RO, -SLCommon.F.FRO, -SLCommon.Redo, -SLCommon.C, -Gconcl_list.BIRO2.IRO, -Gconcl_list.F2.RO, -Gconcl_list.F2.FRO, -Gconcl_list.Simulator, -OSimulator}) + &m : + (forall (F <: ODFUNCTIONALITY) (P <: ODPRIMITIVE), + islossless P.f => + islossless P.fi => + islossless F.f => + islossless Dist(F,P).distinguish) => + `|Pr[OGReal(CSome(Sponge), PSome(Perm), ODRestr(Dist)).main() @ &m : res] - + Pr[OGIdeal(FSome(IRO), OSimulator, ODRestr(Dist)).main() @ &m : res]| <= + (limit ^ 2 - limit)%r / (2 ^ (r + c + 1))%r + (4 * limit ^ 2)%r / (2 ^ c)%r. +proof. +move=>h. +have->: Pr[OGReal(CSome(Sponge), PSome(Perm), ODRestr(Dist)).main() @ &m : res] = + Pr[RealIndif(Sponge, Perm, DRestr(OD(Dist))).main() @ &m : res]. ++ byequiv=>//=; proc; inline*; sim; sp. + call(: ={glob Perm, glob Counter} /\ ={c}(Counter,Cntr))=>/>; auto. + - proc; inline*; sp; auto; if; 1, 3: auto; sp. + by rcondt{2} 1; 1: auto; sp; if; auto. + - proc; inline*; sp; auto; if; auto; sp. + by rcondt{2} 1; 1: auto; sp; if; auto. + proc; inline*; sp; auto; sp; if; auto; sp. + rcondt{2} 1; auto; sp=>/>. + conseq(:_==> ={glob Perm} /\ n{1} = n0{2} /\ z0{1} = z1{2})=> />; sim. + while(={glob Perm, sa, sc, i} /\ (n,z0){1} = (n0,z1){2}); auto. + - by sp; if; auto; sp; if; auto. + conseq(:_==> ={glob Perm, sa, sc})=> />; sim. + by while(={glob Perm, sa, sc, xs}); auto; sp; if; auto=> />. +have->: Pr[OGIdeal(FSome(IRO), OSimulator, ODRestr(Dist)).main() @ &m : res] = + Pr[IdealIndif(IRO, Simulator, DRestr(OD(Dist))).main() @ &m : res]. ++ byequiv=>//=; proc; inline*; sim; sp. + call(: ={glob IRO, glob Simulator, glob Counter} /\ ={c}(Counter,Cntr)); auto. + - proc; inline*; auto; sp; if; auto; sp. + rcondt{2} 1; auto; sp; if; 1, 3: auto; sim; if; 1, 3: auto; sp; sim. + if; [1:by auto=> /> &1 &2 <- /> <- />|3:auto=> />]; sp. + * if; auto=> />. + conseq(:_==> ={IRO.mp} /\ bs0{1} = bs{2})=> />; sim=> />. + by move=> &1 &2 <- /> <- />. + by if; auto=> />; sim=> &1 &2 /> <- /> <- /= ->. + - proc; inline*; sp; auto; if; auto; sp. + by rcondt{2} 1; auto; sp; if; auto. + proc; inline*; sp; auto; if; auto; sp. + rcondt{2} 1; auto; sp; if; auto=> />. + by conseq(:_==> bs{1} = bs0{2} /\ ={IRO.mp, glob Simulator})=> />; sim. +apply (security (OD(Dist)) _ &m). +move=> F P hp hpi hf; proc; inline*; sp. +call (h (OFC(DFSome(F))) (OPC(DPSome(P))) _ _ _); auto. ++ by proc; inline*; sp; if; auto; call hp; auto. ++ by proc; inline*; sp; if; auto; call hpi; auto. +by proc; inline*; sp; if; auto; call hf; auto. +qed. + + + +end section. diff --git a/sha3/proof/SecureHash.eca b/sha3/proof/SecureHash.eca new file mode 100644 index 0000000..db62960 --- /dev/null +++ b/sha3/proof/SecureHash.eca @@ -0,0 +1,147 @@ +require import Int Real SmtMap FSet Distr. +require (****) OptionIndifferentiability. + +type from, to, block. + + +clone import OptionIndifferentiability as OIndif with + type p <- block, + type f_in <- from, + type f_out <- to +proof *. + + +op sampleto : to distr. + +op bound : int. +axiom bound_ge0 : 0 <= bound. + +axiom sampleto_ll: is_lossless sampleto. +axiom sampleto_full: is_full sampleto. +axiom sampleto_fu: is_funiform sampleto. + +(* clone import PROM.GenEager as RO with *) +(* type from <- from, *) +(* type to <- to, *) +(* op sampleto <- fun _ => sampleto *) +(* proof * by exact/sampleto_ll. *) + +op increase_counter (c : int) (m : from) : int. +axiom increase_counter_spec c m : c <= increase_counter c m. + + +(* module type RF = { *) +(* proc init() : unit *) +(* proc get(x : from) : to option *) +(* proc sample (x: from) : unit *) +(* }. *) + +(* module RF (R : RO) : RF = { *) +(* proc init = R.init *) +(* proc get (x : from) : to option = { *) +(* var y; *) +(* y <@ R.get(x); *) +(* return Some y; *) +(* } *) +(* proc sample = R.sample *) +(* }. *) + +module Bounder = { + var bounder : int + proc init () = { + bounder <- 0; + } +}. + +module FBounder (F : OFUNCTIONALITY) : OFUNCTIONALITY = { + proc init () : unit = { + Bounder.init(); + F.init(); + } + proc f(x : from) : to option = { + var y : to option <- None; + if (increase_counter Bounder.bounder x <= bound) { + Bounder.bounder <- increase_counter Bounder.bounder x; + y <@ F.f(x); + } + return y; + } +}. + + +module PBounder (P : OPRIMITIVE) : OPRIMITIVE = { + proc init () = { + P.init(); + Bounder.init(); + } + proc f (x : block) : block option = { + var y <- None; + if (Bounder.bounder < bound) { + y <@ P.f(x); + Bounder.bounder <- Bounder.bounder + 1; + } + return y; + } + proc fi (x : block) : block option = { + var y <- None; + if (Bounder.bounder < bound) { + y <@ P.fi(x); + Bounder.bounder <- Bounder.bounder + 1; + } + return y; + } +}. + +module type AdvPreimage (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + proc guess(h : to) : from +}. + +module Preimage (A : AdvPreimage, F : OCONSTRUCTION, P : OPRIMITIVE) = { + proc main () : bool = { + var m,hash,hash'; + hash <$ sampleto; + PBounder(P).init(); + FBounder(F(P)).init(); + m <@ A(FBounder(F(P)),PBounder(P)).guess(hash); + hash' <@ FBounder(F(P)).f(m); + return hash' = Some hash; + } +}. + + +(*-------------------------------------------------------------------------*) +module type AdvSecondPreimage (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + proc guess(m : from) : from +}. + +module SecondPreimage (A : AdvSecondPreimage, F : OCONSTRUCTION, P : OPRIMITIVE) = { + proc main (m1 : from) : bool = { + var m2, hash1, hash2; + PBounder(P).init(); + FBounder(F(P)).init(); + m2 <@ A(FBounder(F(P)),PBounder(P)).guess(m1); + hash1 <@ FBounder(F(P)).f(m1); + hash2 <@ FBounder(F(P)).f(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } +}. + + +(*--------------------------------------------------------------------------*) +module type AdvCollision (F : ODFUNCTIONALITY) (P : ODPRIMITIVE) = { + proc guess() : from * from +}. + +module Collision (A : AdvCollision, F : OCONSTRUCTION, P : OPRIMITIVE) = { + proc main () : bool = { + var m1,m2,hash1,hash2; + PBounder(P).init(); + FBounder(F(P)).init(); + (m1,m2) <@ A(FBounder(F(P)),PBounder(P)).guess(); + hash1 <@ FBounder(F(P)).f(m1); + hash2 <@ FBounder(F(P)).f(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } +}. + + diff --git a/sha3/proof/SecureORO.eca b/sha3/proof/SecureORO.eca new file mode 100644 index 0000000..ee3a014 --- /dev/null +++ b/sha3/proof/SecureORO.eca @@ -0,0 +1,492 @@ +require import Int Distr Real SmtMap FSet Mu_mem. +require (*--*) StdOrder. +(*---*) import StdOrder.IntOrder. +require (****) PROM FelTactic. + + +type from, to. + +op sampleto : to distr. + +op bound : int. +axiom bound_ge0 : 0 <= bound. + +axiom sampleto_ll: is_lossless sampleto. +axiom sampleto_full: is_full sampleto. +axiom sampleto_fu: is_funiform sampleto. + +clone import PROM.FullRO as RO with + type in_t <- from, + type out_t <- to, + op dout _ <- sampleto, + type d_in_t <- unit, + type d_out_t <- bool. +import FullEager. + +op increase_counter (c : int) (m : from) : int. +axiom increase_counter_spec c m : c <= increase_counter c m. + +module type RF = { + proc init() : unit + proc get(x : from) : to option + proc sample (x: from) : unit +}. + +module RF (R : RO) : RF = { + proc init = R.init + proc get (x : from) : to option = { + var y; + y <@ R.get(x); + return Some y; + } + proc sample = R.sample +}. + +module Bounder (F : RF) : RF = { + var bounder : int + proc init () : unit = { + bounder <- 0; + F.init(); + } + proc get(x : from) : to option = { + var y : to option <- None; + if (bounder < bound) { + bounder <- bounder + 1; + y <@ F.get(x); + } + return y; + } + proc sample = F.sample +}. + + +module type Oracle = { + proc get(x : from) : to option {} +}. + +module type AdvPreimage (F : Oracle) = { + proc guess(h : to) : from +}. + +module Preimage (A : AdvPreimage, F : RF) = { + proc main () : bool = { + var m,hash,hash'; + hash <$ sampleto; + Bounder(F).init(); + m <@ A(Bounder(F)).guess(hash); + hash' <@ Bounder(F).get(m); + return hash' = Some hash; + } +}. + +section Preimage. + + declare module A <: AdvPreimage {-RO, -Preimage}. + + local module FEL (A : AdvPreimage, F : RF) = { + proc main (hash : to) : from = { + var m; + Bounder(F).init(); + m <@ A(Bounder(F)).guess(hash); + return m; + } + }. + + local module Preimage2 (A : AdvPreimage, F : RF) = { + var hash : to + proc main () : bool = { + var m,hash'; + hash <$ sampleto; + m <@ FEL(A,F).main(hash); + hash' <@ Bounder(F).get(m); + return hash' = Some hash; + } + }. + + lemma RO_is_preimage_resistant &m : + Pr [ Preimage(A,RF(RO)).main() @ &m : res ] <= (bound + 1)%r * mu1 sampleto witness. + proof. + have->: Pr [ Preimage (A,RF(RO)).main() @ &m : res ] = + Pr [ Preimage2(A,RF(RO)).main() @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + byphoare(: _ ==> _) => //=; proc. + seq 2 : (rng RO.m Preimage2.hash) (bound%r * mu1 sampleto witness) 1%r 1%r + (mu1 sampleto witness) + (card (fdom RO.m) <= Bounder.bounder <= bound). + + inline*; auto; call(: card (fdom RO.m) <= Bounder.bounder <= bound)=> //=. + - proc; inline*; auto; sp; if; 2:auto; wp. + wp; conseq(:_==> card (fdom RO.m) + 1 <= Bounder.bounder <= bound); 2: by auto;smt(). + move=> &h /> c H1 H2 c2 r x h1 h2; split; 2: smt(). + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + by auto=> />; rewrite fdom0 fcards0; smt(bound_ge0). + + seq 1 : true 1%r (bound%r * mu1 sampleto witness) 0%r _; auto. + exists * Preimage2.hash; elim* => h. + call(: Preimage2.hash = h /\ h = arg ==> rng RO.m h)=> //. + bypr=> /> {&m} &m <<- <-. + pose h := Preimage2.hash{m}. + have H: forall &m h, + Pr[FEL(A, RF(RO)).main(h) @ &m : rng RO.m h] <= bound%r * mu1 sampleto witness; last first. + + exact(H &m h). + move=> {&m h} &m h. + fel 1 Bounder.bounder (fun _, mu1 sampleto witness) bound (rng RO.m h) + [Bounder(RF(RO)).get: (card (fdom RO.m) <= Bounder.bounder < bound)] + (card (fdom RO.m) <= Bounder.bounder <= bound) + =>//. + - rewrite StdBigop.Bigreal.BRA.big_const List.count_predT List.Range.size_range. + rewrite ler_maxr //=; 1:smt(bound_ge0). + rewrite-RField.AddMonoid.iteropE-RField.intmulpE; 1: smt(bound_ge0). + by rewrite RField.intmulr; smt(). + - inline*; auto=> />. + rewrite mem_rng_empty /= fdom0 fcards0 /=; smt(bound_ge0). + - proc. + sp; if; auto; sp; inline*; sp; wp=> /=. + case: (x \in RO.m); wp => //=. + + by hoare; auto; smt(mu_bounded). + rnd (pred1 h); auto=> /> &h c H0 H1 H2 H3 H4 H5. + rewrite (sampleto_fu h witness) /= => ? ?. + rewrite rngE/= => [][] a; rewrite get_setE. + case: (a=x{h}) => [->>|] //=. + by move:H2; rewrite rngE /= negb_exists/= => /(_ a) //=. + - move=> c; proc; inline*; sp; if; sp. + + auto; progress. + + smt(). + + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + + smt(). + + smt(). + + smt(). + + smt(). + by auto. + move=> b c; proc; sp; inline*; if; auto; progress. + - rewrite 2!rngE /= eq_iff; split=> [][] a. + + by rewrite get_setE; move: H4; rewrite domE /=; smt(). + move=> H8; exists a; rewrite get_setE; move: H4; rewrite domE /=; smt(). + - smt(). + - by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + - smt(). + - smt(). + - smt(). + - smt(). + + by inline*; auto. + + by inline*; auto. + + inline*; sp; wp. + if; sp; wp; last by hoare;auto;progress; smt(mu_bounded). + case: (x \in RO.m). + - hoare; auto; progress. + rewrite H3/=; move: H1; rewrite rngE /= negb_exists /=. + by have:=H3; rewrite domE; smt(). + rnd (pred1 Preimage2.hash); auto=> /> &hr *. + rewrite (sampleto_fu Preimage2.hash{hr} witness)/= => ??. + by rewrite get_setE /=; smt(). + smt(). + qed. + +end section Preimage. + +(*-------------------------------------------------------------------------*) +module type AdvSecondPreimage (F : Oracle) = { + proc guess(m : from) : from +}. + +module SecondPreimage (A : AdvSecondPreimage, F : RF) = { + proc main (m1 : from) : bool = { + var m2, hash1, hash2; + Bounder(F).init(); + m2 <@ A(Bounder(F)).guess(m1); + hash1 <@ Bounder(F).get(m1); + hash2 <@ Bounder(F).get(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } +}. + +section SecondPreimage. + + + declare module A <: AdvSecondPreimage {-Bounder, -RO, -FRO}. + + local module FEL (A : AdvSecondPreimage, F : RF) = { + proc main (m1 : from) : from = { + var m2; + Bounder(F).init(); + Bounder(F).sample(m1); + m2 <@ A(Bounder(F)).guess(m1); + return m2; + } + }. + + local module SecondPreimage2 (A : AdvSecondPreimage, F : RF) = { + var m2 : from + proc main (m1 : from) : bool = { + var hash1,hash2; + m2 <@ FEL(A,F).main(m1); + hash1 <@ Bounder(F).get(m1); + hash2 <@ Bounder(F).get(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } + }. + + local module D1 (A : AdvSecondPreimage, F : RO) = { + var m1 : from + proc distinguish () : bool = { + var b; + b <@ SecondPreimage2(A,RF(F)).main(m1); + return b; + } + }. + + local module SecondPreimage3 (A : AdvSecondPreimage, F : RO) = { + proc main (m1 : from) : bool = { + var b; + SecondPreimage2.m2 <- witness; + D1.m1 <- m1; + Bounder(RF(F)).init(); + b <@ D1(A,F).distinguish(); + return b; + } + }. + + + lemma RO_is_second_preimage_resistant &m (mess1 : from) : + Pr [ SecondPreimage(A,RF(RO)).main(mess1) @ &m : res ] + <= (bound + 1)%r * mu1 sampleto witness. + proof. + have->: Pr [ SecondPreimage(A,RF(RO)).main(mess1) @ &m : res ] = + Pr [ SecondPreimage(A,RF(LRO)).main(mess1) @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + have->: Pr [ SecondPreimage(A,RF(LRO)).main(mess1) @ &m : res ] = + Pr [ SecondPreimage2(A,RF(LRO)).main(mess1) @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + have->: Pr [ SecondPreimage2(A,RF(LRO)).main(mess1) @ &m : res ] = + Pr [ SecondPreimage2(A,RF(RO)).main(mess1) @ &m : res ]. + + have->: Pr [ SecondPreimage2(A,RF(LRO)).main(mess1) @ &m : res ] = + Pr [ SecondPreimage3(A,LRO).main(mess1) @ &m : res ]. + - by byequiv=> //=; proc; inline*; wp 15 -2; sim. + have->: Pr [ SecondPreimage3(A,LRO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage3(A,RO).main(mess1) @ &m : res ]. + - rewrite eq_sym. + byequiv=>//=; proc. + call(RO_LRO_D (D1(A)) _); first by rewrite sampleto_ll. + by inline*; auto. + by byequiv=> //=; proc; inline*; wp -2 18; sim. + byphoare(: arg = mess1 ==> _)=>//=; proc. + seq 1 : (rng (rem RO.m mess1) (oget RO.m.[mess1])) + (bound%r * mu1 sampleto witness) 1%r + 1%r (mu1 sampleto witness) + (card (fdom RO.m) - 1 <= Bounder.bounder <= bound + /\ mess1 \in RO.m /\ mess1 = m1). + + inline*; auto; call(: card (fdom RO.m) - 1 <= Bounder.bounder <= bound + /\ mess1 \in RO.m). + - proc; inline*; auto; sp; if; last by auto; smt(). + auto=> /> &h c Hc Hdom Hc2 sample. + by rewrite sampleto_full/=!fdom_set !fcardU !fcard1;smt(mem_set fcard_ge0). + auto=> /> sample. + by rewrite mem_set mem_empty/= fdom_set fdom0 fset0U fcard1; smt(bound_ge0). + + call(: arg = mess1 ==> rng (rem RO.m mess1) (oget RO.m.[mess1])); auto. + bypr=> {&m} &m h; rewrite h. + fel 2 Bounder.bounder (fun _, mu1 sampleto witness) bound + (mess1 \in RO.m /\ rng (rem RO.m mess1) (oget RO.m.[mess1])) + [Bounder(RF(RO)).get: (card (fdom RO.m) - 1 <= Bounder.bounder < bound)] + (card (fdom RO.m) - 1 <= Bounder.bounder <= bound /\ mess1 \in RO.m)=> {h} + =>//. + + rewrite StdBigop.Bigreal.BRA.big_const List.count_predT List.Range.size_range. + rewrite ler_maxr //=; 1:smt(bound_ge0). + rewrite-RField.AddMonoid.iteropE-RField.intmulpE; 1: smt(bound_ge0). + by rewrite RField.intmulr; smt(mu_bounded bound_ge0). + + inline*; auto=> />. + move=> r; rewrite mem_empty /= !mem_set mem_empty/= sampleto_full /=. + rewrite get_set_sameE//= fdom_set fdom0 fset0U fcard1 /= rngE /=; split; 2: smt(bound_ge0). + by rewrite negb_exists/= => a; rewrite remE get_setE //= emptyE; smt(). + + proc; inline*; sp; if; last by hoare; auto. + sp; case: (x \in RO.m)=> //=. + - by hoare; auto; smt(mu_bounded). + rcondt 2; 1: auto; wp=> /=. + conseq(:_ ==> pred1 (oget RO.m.[mess1]) r)=> />. + - move=> /> &h c H0c Hcb Hnrng Hmc _ Hdom1 Hdom2 sample. + rewrite mem_set Hdom1 /= get_set_neqE; 1: smt(). + have->: (rem RO.m{h}.[x{h} <- sample] mess1) = (rem RO.m{h} mess1).[x{h} <- sample]. + + by apply fmap_eqP=> y; rewrite remE 2!get_setE remE; smt(). + move: Hnrng; rewrite Hdom1 /= rngE /= negb_exists /= => Hnrng. + rewrite rngE/= => [][] mess; rewrite get_setE remE. + by have:= Hnrng mess; rewrite remE; smt(). + rnd; auto; progress. + by have ->:= sampleto_fu witness (oget RO.m{hr}.[mess1]). + + move=> c; proc; inline*; sp; if; auto; progress. + - smt(). + - by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + - smt(). + - smt(mem_set). + - smt(). + - smt(). + - smt(). + + move=> b c; proc; inline*; sp; if; auto; smt(). + + by inline*; auto. + + by auto. + + inline*; sp; auto. + if; sp; last first. + + sp; hoare; auto; if; auto. + case(Bounder.bounder < bound); last first. + - by rcondf 8; 1: auto; hoare; auto. + rcondt 8; 1: auto. + swap 11 -8; sp. + swap [7..11] -6; sp. + swap[5..6] 2; wp 6=> /=. + case: (SecondPreimage2.m2 \in RO.m). + - rcondf 5; 1: auto; hoare; auto=> /> &h d _ _ in_dom1 not_rng d_bound _ in_dom2. + move=> sample2 _ m1_in_RO sample1 _; rewrite negb_and/=. + move: not_rng; rewrite rngE /= negb_exists /= => /(_ SecondPreimage2.m2{h}). + rewrite remE; case: (SecondPreimage2.m2{h} = m1{h})=> //=. + by move: in_dom1 in_dom2; smt(). + rcondt 5; 1: auto; wp. + rnd (pred1 (oget RO.m.[x3])); auto. + move => /> &h d _ _ in_dom1 not_rng _ _ nin_dom2 sample2 _. + rewrite (sampleto_fu (oget RO.m{h}.[m1{h}]) witness) /= => _ sample1 _. + by rewrite get_set_sameE//=; smt(). + smt(). + qed. + +end section SecondPreimage. + + +(*--------------------------------------------------------------------------*) +module type AdvCollision (F : Oracle) = { + proc guess() : from * from +}. + +module Collision (A : AdvCollision, F : RF) = { + proc main () : bool = { + var m1,m2,hash1,hash2; + Bounder(F).init(); + (m1,m2) <@ A(Bounder(F)).guess(); + hash1 <@ Bounder(F).get(m1); + hash2 <@ Bounder(F).get(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } +}. + +section Collision. + + declare module A <: AdvCollision {-RO, -FRO, -Bounder}. + + local module FEL (A : AdvCollision, F : RF) = { + proc main () : from * from = { + var m1,m2; + Bounder(F).init(); + (m1,m2) <@ A(Bounder(F)).guess(); + return (m1,m2); + } + }. + + local module Collision2 (A : AdvCollision) (F : RF) = { + proc main () : bool = { + var m1,m2,hash1,hash2; + (m1,m2) <@ FEL(A,F).main(); + hash1 <@ Bounder(F).get(m1); + hash2 <@ Bounder(F).get(m2); + return m1 <> m2 /\ exists y, Some y = hash1 /\ Some y = hash2; + } + }. + + op collision (m : ('a, 'b) fmap) = + exists m1 m2, m1 <> m2 /\ m1 \in m /\ m2 \in m /\ m.[m1] = m.[m2]. + + lemma RO_is_collision_resistant &m : + Pr [ Collision(A,RF(RO)).main() @ &m : res ] + <= ((bound * (bound - 1) + 2)%r / 2%r * mu1 sampleto witness). + proof. + have->: Pr [ Collision(A,RF(RO)).main() @ &m : res ] = + Pr [ Collision2(A,RF(RO)).main() @ &m : res ]. + + by byequiv=>//=; proc; inline*; sim. + byphoare=> //; proc. + seq 1 : (collision RO.m) + ((bound * (bound - 1))%r / 2%r * mu1 sampleto witness) 1%r + 1%r (mu1 sampleto witness) + (card (fdom RO.m) <= Bounder.bounder <= bound); first last; first last. + + auto. + + auto. + + inline*; sp. + if; sp; last first. + - by wp; conseq(:_==> false)=> />; hoare; auto. + case: (Bounder.bounder < bound); last first. + - rcondf 8; 1:auto; hoare; auto. + rcondt 8; 1: auto. + swap 11 -8. + swap [7..11] -6; sp. + swap [5..6] 1; wp 5=> /=. + swap 3 -1. + case: (m1 = m2). + - by hoare; auto. + case: (m1 \in RO.m); case: (m2 \in RO.m). + - rcondf 3; 1: auto; rcondf 4; 1: auto; hoare; auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 _ _ _ _. + move: Hcoll; rewrite /collision negb_exists /= => /(_ m1{h}). + rewrite negb_exists /= => /(_ m2{h}). + by rewrite neq12 in_dom1 in_dom2 /=; smt(). + - rcondf 3; 1: auto; rcondt 4; 1: auto; wp. + rnd (pred1 (oget RO.m.[x3])). + auto=> /> &h d Hmc Hcb Hcoll _ _ neq12 in_dom1 in_dom2 _ _; split. + * smt(sampleto_fu). + by move=> _ sample _; rewrite get_set_sameE; smt(). + - rcondt 3; 1: auto; rcondf 5; 1: (by auto; smt(mem_set)). + swap 1; wp=> /=; rnd (pred1 (oget RO.m.[m2])); auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 _ _; split. + * smt(sampleto_fu). + move=> _ sample _. + by rewrite get_set_sameE get_set_neqE 1:eq_sym. + rcondt 3; 1: auto; rcondt 5; 1: (by auto; smt(mem_set)). + swap 2 -1; swap 1; wp=> //=; rnd (pred1 r); auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 sample1 _; split. + * smt(sampleto_fu). + move=> _ sample2 _. + by rewrite get_set_sameE get_set_sameE; smt(). + + by move=> />; smt(mu_bounded). + + inline*; wp; call(: card (fdom RO.m) <= Bounder.bounder <= bound); auto. + - proc; inline*; sp; if; last by auto; smt(). + auto=> /> &h d Hbc Hcb sample _; split. + * by move=> nin_dom1; rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + by move=> in_dom1; smt(). + by move=> />; rewrite fdom0 fcards0; smt(bound_ge0). + call(: true ==> collision RO.m); auto; bypr=> /> {&m} &m. + fel 1 Bounder.bounder (fun i, i%r * mu1 sampleto witness) bound + (collision RO.m) + [Bounder(RF(RO)).get: (card (fdom RO.m) <= Bounder.bounder < bound)] + (card (fdom RO.m) <= Bounder.bounder <= bound)=> //. + + rewrite -StdBigop.Bigreal.BRA.mulr_suml RField.mulrAC. + rewrite StdOrder.RealOrder.ler_wpmul2r; 1: smt(mu_bounded). + by rewrite StdBigop.Bigreal.sumidE //; smt(bound_ge0). + + inline*; auto=> />. + rewrite fdom0 fcards0; split; 2: smt(bound_ge0). + rewrite /collision negb_exists /= => a; rewrite negb_exists /= => b. + by rewrite mem_empty. + + bypr=> /> {&m} &m; pose c := Bounder.bounder{m}; move=> H0c Hcbound Hcoll Hmc _. + byphoare(: !collision RO.m /\ card (fdom RO.m) <= c ==> _)=>//=. + proc; inline*; sp; if; last first. + - by hoare; auto; smt(mu_bounded). + case: (x \in RO.m). + - by hoare; auto; smt(mu_bounded). + rcondt 5; 1: auto; sp; wp=> /=. + conseq(:_==> r \in frng RO.m). + - move=> /> &h c2 Hcoll2 Hb2c Hc2b nin_dom sample m1 m2 neq. + rewrite 2!mem_set. + case: (m1 = x{h}) => //=. + * move=> <<-; rewrite eq_sym neq /= get_set_sameE get_set_neqE//= 1:eq_sym //. + by rewrite mem_frng rngE /= => _ ->; exists m2. + case: (m2 = x{h}) => //=. + * move=> <<- _ in_dom1. + by rewrite get_set_neqE // get_set_sameE mem_frng rngE/= => <-; exists m1. + move=> neq2 neq1 in_dom1 in_dom2; rewrite get_set_neqE // get_set_neqE //. + have:= Hcoll2; rewrite negb_exists /= => /(_ m1). + rewrite negb_exists /= => /(_ m2). + by rewrite neq in_dom1 in_dom2 /= => ->. + rnd; skip=> /> &h bounder _ h _. + rewrite (mu_mem (frng RO.m{h}) sampleto (mu1 sampleto witness)); 1: smt(sampleto_fu). + rewrite StdOrder.RealOrder.ler_wpmul2r //. + by rewrite le_fromint; smt(le_card_frng_fdom). + + move=> c; proc; inline*; auto; sp; if; last by auto; smt(). + auto=> /> &h h1 h2 _ sample _. + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + move=> b c; proc; inline*; sp; if; auto. + move=> /> &h h1 h2 _ h3 sample _. + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + qed. + + +end section Collision. diff --git a/sha3/proof/SecureRO.eca b/sha3/proof/SecureRO.eca new file mode 100644 index 0000000..c7df62a --- /dev/null +++ b/sha3/proof/SecureRO.eca @@ -0,0 +1,525 @@ +require import Int Distr Real SmtMap FSet Mu_mem. +require (*--*) StdOrder. +(*---*) import StdOrder.IntOrder. +require (****) PROM FelTactic. + + +type from, to. + +op sampleto : to distr. + +op bound : int. +axiom bound_gt0 : 0 < bound. + +axiom sampleto_ll: is_lossless sampleto. +axiom sampleto_full: is_full sampleto. +axiom sampleto_fu: is_funiform sampleto. + +clone import PROM.FullRO as RO with + type in_t <- from, + type out_t <- to, + op dout _ <- sampleto, + type d_in_t <- unit, + type d_out_t <- bool. +import FullEager. + +op increase_counter (c : int) (m : from) : int. +axiom increase_counter_spec c m : c <= increase_counter c m. + +op bound_counter : int. +axiom bound_counter_ge0 : 0 <= bound_counter. + +module Counter = { + var c : int + proc init() = { + c <- 0; + } +}. + +module type RF = { + proc init() : unit + proc get(x : from) : to +}. + +module Bounder (F : RF) : RF = { + var bounder : int + proc init () : unit = { + bounder <- 0; + Counter.init(); + F.init(); + } + proc get(x : from) : to = { + var y : to <- witness; + if (bounder < bound) { + bounder <- bounder + 1; + if (increase_counter Counter.c x < bound_counter) { + Counter.c <- increase_counter Counter.c x; + y <@ F.get(x); + } + } + return y; + } +}. + + +module type Oracle = { + proc get(x : from) : to {} +}. + +module type AdvPreimage (F : Oracle) = { + proc guess(h : to) : from +}. + +module Preimage (A : AdvPreimage, F : RF) = { + proc main (hash : to) : bool = { + var b,m,hash'; + Counter.init(); + Bounder(F).init(); + m <@ A(Bounder(F)).guess(hash); + if (increase_counter Counter.c m < bound_counter) { + hash' <@ F.get(m); + b <- hash = hash'; + } else b <- false; + return b; + } +}. + +section Preimage. + + declare module A <: AdvPreimage {-RO, -Preimage}. + + local module FEL (A : AdvPreimage, F : RF) = { + proc main (hash : to) : from = { + var m; + Bounder(F).init(); + m <@ A(Bounder(F)).guess(hash); + return m; + } + }. + + local module Preimage2 (A : AdvPreimage, F : RF) = { + proc main (hash : to) : bool = { + var b,m,hash'; + m <@ FEL(A,F).main(hash); + if (increase_counter Counter.c m < bound_counter) { + hash' <@ F.get(m); + b <- hash = hash'; + } else b <- false; + return b; + } + }. + + lemma RO_is_preimage_resistant &m (h : to) : + Pr [ Preimage(A,RO).main(h) @ &m : res ] <= (bound + 1)%r * mu1 sampleto h. + proof. + have->: Pr [ Preimage (A,RO).main(h) @ &m : res ] = + Pr [ Preimage2(A,RO).main(h) @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + byphoare(: arg = h ==> _) => //=; proc. + seq 1 : (rng RO.m h) (bound%r * mu1 sampleto h) 1%r 1%r (mu1 sampleto h) + (card (fdom RO.m) <= Bounder.bounder <= bound /\ hash = h). + + inline*; auto; call(: card (fdom RO.m) <= Bounder.bounder <= bound)=> //=. + - proc; inline*; auto; sp; if; 2:auto; wp; sp. + if; last by auto; smt(). + wp; conseq(:_==> card (fdom RO.m) + 1 <= Bounder.bounder <= bound); 2: by auto;smt(). + move=> &h /> c H1 _ H2 c2 r x h1 h2; split; 2: smt(). + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + by auto=> />; rewrite fdom0 fcards0; smt(bound_gt0). + + call(: true ==> rng RO.m h)=> //; bypr=> /> {&m} &m. + fel 1 Bounder.bounder (fun _, mu1 sampleto h) bound (rng RO.m h) + [Bounder(RO).get: (card (fdom RO.m) <= Bounder.bounder < bound)] + (card (fdom RO.m) <= Bounder.bounder <= bound) + =>//. + - rewrite StdBigop.Bigreal.BRA.big_const List.count_predT List.Range.size_range. + rewrite ler_maxr //=; 1:smt(bound_gt0). + rewrite-RField.AddMonoid.iteropE-RField.intmulpE; 1: smt(bound_gt0). + by rewrite RField.intmulr; smt(). + - inline*; auto=> />. + by rewrite mem_rng_empty /= fdom0 fcards0 /=; smt(bound_gt0). + - proc. + sp; if; auto; sp; inline*; sp; wp=> /=. + if; last by hoare; auto; progress; smt(mu_bounded). + case: (x \in RO.m); wp => //=. + + by hoare; auto; smt(mu_bounded). + rnd (pred1 h); auto=> /> &h c ge0_c lt_c_bound h_notin_rngRO le_sRO_c le_c_bound Hcount x_notin_RO v _. + rewrite rngE=> /= - [] a; rewrite get_setE. + case: (a=x{h}) => [->>|] //=. + by move:h_notin_rngRO; rewrite rngE /= negb_exists/= => /(_ a) //=. + - move=> c; proc; inline*; sp; if; sp. + + if; auto; progress. + + smt(). + + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + + smt(). + + smt(). + + smt(). + + smt(). + + smt(). + + smt(). + + smt(). + by auto. + move=> b c; proc; sp; if; auto; inline*; auto; sp; if; auto; progress. + - rewrite 2!rngE /= eq_iff; split=> [][] a. + + by rewrite get_setE; move: H5; rewrite domE /=; smt(). + move=> H8; exists a; rewrite get_setE; move: H5; rewrite domE /=; smt(). + - smt(). + - by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + - smt(). + - smt(). + - smt(). + - smt(). + - smt(). + - smt(). + smt(). + + by inline*; auto. + + by inline*; auto. + + inline*; sp; wp. + if; sp; wp; last by hoare;auto;progress; smt(mu_bounded). + case: (x \in RO.m). + - hoare; auto; progress. + rewrite H3/=; move: H1; rewrite rngE /= negb_exists /=. + by have:=H3; rewrite domE; smt(). + rnd (pred1 h); auto=> //= &hr [#]->>H0 H1<<-H2 H3 H4 H5 H6. + by rewrite H4 /= get_setE /=; smt(). + smt(). + qed. + +end section Preimage. + +(*-------------------------------------------------------------------------*) +module type AdvSecondPreimage (F : Oracle) = { + proc guess(m : from) : from +}. + +module SecondPreimage (A : AdvSecondPreimage, F : RF) = { + proc main (m1 : from) : bool = { + var b, m2, hash1, hash2; + Bounder(F).init(); + m2 <@ A(Bounder(F)).guess(m1); + if (increase_counter Counter.c m1 < bound_counter) { + Counter.c <- increase_counter Counter.c m1; + hash1 <@ F.get(m1); + if (increase_counter Counter.c m2 < bound_counter) { + Counter.c <- increase_counter Counter.c m2; + hash2 <@ F.get(m2); + b <- hash1 = hash2 /\ m1 <> m2; + } else b <- false; + } + else b <- false; + return b; + } +}. + +section SecondPreimage. + + + declare module A <: AdvSecondPreimage {-Bounder, -RO, -FRO}. + + local module FEL (A : AdvSecondPreimage, F : RO) = { + proc main (m1 : from) : from = { + var m2; + Bounder(F).init(); + F.sample(m1); + m2 <@ A(Bounder(F)).guess(m1); + return m2; + } + }. + + local module SecondPreimage2 (A : AdvSecondPreimage, F : RO) = { + var m2 : from + proc main (m1 : from) : bool = { + var b, hash1,hash2; + m2 <@ FEL(A,F).main(m1); + if (increase_counter Counter.c m1 < bound_counter) { + Counter.c <- increase_counter Counter.c m1; + hash1 <@ F.get(m1); + if (increase_counter Counter.c m2 < bound_counter) { + Counter.c <- increase_counter Counter.c m2; + hash2 <@ F.get(m2); + b <- hash1 = hash2 /\ m1 <> m2; + } else b <- false; + } + else b <- false; + return b; + } + }. + + local module D1 (A : AdvSecondPreimage, F : RO) = { + var m1 : from + proc distinguish () : bool = { + var b; + b <@ SecondPreimage2(A,F).main(m1); + return b; + } + }. + + local module SecondPreimage3 (A : AdvSecondPreimage, F : RO) = { + proc main (m1 : from) : bool = { + var b; + SecondPreimage2.m2 <- witness; + D1.m1 <- m1; + Bounder(F).init(); + b <@ D1(A,F).distinguish(); + return b; + } + }. + + + lemma RO_is_second_preimage_resistant &m (mess1 : from) : + Pr [ SecondPreimage(A,RO).main(mess1) @ &m : res ] + <= (bound + 1)%r * mu1 sampleto witness. + proof. + have->: Pr [ SecondPreimage(A,RO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage(A,LRO).main(mess1) @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + have->: Pr [ SecondPreimage(A,LRO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage2(A,LRO).main(mess1) @ &m : res ]. + + by byequiv=> //=; proc; inline*; sim. + have->: Pr [ SecondPreimage2(A,LRO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage2(A,RO).main(mess1) @ &m : res ]. + + have->: Pr [ SecondPreimage2(A,LRO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage3(A,LRO).main(mess1) @ &m : res ]. + - by byequiv=> //=; proc; inline*; wp 15 -2; sim. + have->: Pr [ SecondPreimage3(A,LRO).main(mess1) @ &m : res ] = + Pr [ SecondPreimage3(A,RO).main(mess1) @ &m : res ]. + - rewrite eq_sym. + byequiv=>//=; proc. + call(RO_LRO_D (D1(A)) _); first by rewrite sampleto_ll. + by inline*; auto. + by byequiv=> //=; proc; inline*; wp -2 18; sim. + byphoare(: arg = mess1 ==> _)=>//=; proc. + seq 1 : (rng (rem RO.m mess1) (oget RO.m.[mess1])) + (bound%r * mu1 sampleto witness) 1%r + 1%r (mu1 sampleto witness) + (card (fdom RO.m) - 1 <= Bounder.bounder <= bound + /\ mess1 \in RO.m /\ mess1 = m1). + + inline*; auto; call(: card (fdom RO.m) - 1 <= Bounder.bounder <= bound + /\ mess1 \in RO.m). + - proc; inline*; auto; sp; if; sp; auto; if; last by auto; smt(). + auto=> /> &h c Hc _ Hdom Hc2 _ sample. + by rewrite sampleto_full/=!fdom_set !fcardU !fcard1;smt(mem_set fcard_ge0). + auto=> /> sample. + by rewrite mem_set mem_empty/= fdom_set fdom0 fset0U fcard1; smt(bound_gt0). + + call(: arg = mess1 ==> rng (rem RO.m mess1) (oget RO.m.[mess1])); auto. + bypr=> {&m} &m h; rewrite h. + fel 2 Bounder.bounder (fun _, mu1 sampleto witness) bound + (mess1 \in RO.m /\ rng (rem RO.m mess1) (oget RO.m.[mess1])) + [Bounder(RO).get: (card (fdom RO.m) - 1 <= Bounder.bounder < bound)] + (card (fdom RO.m) - 1 <= Bounder.bounder <= bound /\ mess1 \in RO.m)=> {h} + =>//. + + rewrite StdBigop.Bigreal.BRA.big_const List.count_predT List.Range.size_range. + rewrite ler_maxr //=; 1:smt(bound_gt0). + rewrite-RField.AddMonoid.iteropE-RField.intmulpE; 1: smt(bound_gt0). + by rewrite RField.intmulr; smt(mu_bounded bound_gt0). + + inline*; auto=> /> r. + rewrite mem_empty /= !mem_set mem_empty/= sampleto_full /=. + rewrite get_set_sameE//= fdom_set fdom0 fset0U fcard1 /= rngE /=; split; 2: smt(bound_gt0). + by rewrite negb_exists/= => a; rewrite remE get_setE //= emptyE; smt(). + + proc; inline*; sp; if; last by hoare; auto. + sp; if; sp; last by hoare; auto; smt(mu_bounded). + case: (x0 \in RO.m)=> //=. + - by hoare; auto; smt(mu_bounded). + rcondt 2; 1: auto; wp=> /=. + conseq(:_ ==> pred1 (oget RO.m.[mess1]) r)=> />. + - move=> /> &h d c H0c Hcb Hnrng Hmc _ Hdom1 _ Hdom2 sample. + rewrite mem_set Hdom1 /= get_set_neqE; 1: smt(). + have->: (rem RO.m{h}.[x{h} <- sample] mess1) = (rem RO.m{h} mess1).[x{h} <- sample]. + + by apply fmap_eqP=> y; rewrite remE 2!get_setE remE; smt(). + move: Hnrng; rewrite Hdom1 /= rngE /= negb_exists /= => Hnrng. + rewrite rngE/= => [][] mess; rewrite get_setE remE. + by have:= Hnrng mess; rewrite remE; smt(). + rnd; auto; progress. + by have ->:= sampleto_fu witness (oget RO.m{hr}.[mess1]). + + move=> c; proc; inline*; sp; if; auto; sp; if; auto; progress. + - smt(). + - by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + - smt(). + - smt(mem_set). + - smt(). + - smt(). + - smt(). + - smt(). + - smt(). + - smt(). + + move=> b c; proc; inline*; sp; if; auto; sp; if; auto; smt(). + + by inline*; auto. + + by auto. + + inline*. + if; sp; last by hoare; auto; smt(mu_bounded). + rcondf 2; 1: auto. + case(increase_counter Counter.c SecondPreimage2.m2 < bound_counter); last first. + - by rcondf 3; 1: auto; hoare; auto; smt(mu_bounded). + rcondt 3; 1: auto. + swap 3 -2; sp. + case: (SecondPreimage2.m2 \in RO.m). + - rcondf 5; 1: auto; hoare; auto=> /> &h d _ _ in_dom1 not_rng _ _ in_dom2. + move=> sample2 _ sample1 _; rewrite negb_and/=. + move: not_rng; rewrite rngE /= negb_exists /= => /(_ SecondPreimage2.m2{h}). + rewrite remE; case: (SecondPreimage2.m2{h} = m1{h})=> //=. + by move: in_dom1 in_dom2; smt(). + rcondt 5; 1: auto; wp; rnd (pred1 hash1); auto. + move => /> &h d _ _ in_dom1 not_rng _ _ nin_dom2 sample2 _. + rewrite (sampleto_fu (oget RO.m{h}.[m1{h}]) witness) /= => sample1 _. + by rewrite get_set_sameE => ->. + smt(). + qed. + +end section SecondPreimage. + + +(*--------------------------------------------------------------------------*) +module type AdvCollision (F : Oracle) = { + proc guess() : from * from +}. + +module Collision (A : AdvCollision, F : RO) = { + proc main () : bool = { + var b,m1,m2,hash1,hash2; + Bounder(F).init(); + (m1,m2) <@ A(Bounder(F)).guess(); + if (increase_counter Counter.c m1 < bound_counter) { + Counter.c <- increase_counter Counter.c m1; + hash1 <@ F.get(m1); + if (increase_counter Counter.c m2 < bound_counter) { + Counter.c <- increase_counter Counter.c m2; + hash2 <@ F.get(m2); + b <- hash1 = hash2 /\ m1 <> m2; + } else b <- false; + } + else b <- false; + return b; + } +}. + +section Collision. + + declare module A <: AdvCollision {-RO, -FRO, -Bounder}. + + local module FEL (A : AdvCollision, F : RO) = { + proc main () : from * from = { + var m1,m2; + Bounder(F).init(); + (m1,m2) <@ A(Bounder(F)).guess(); + return (m1,m2); + } + }. + + local module Collision2 (A : AdvCollision) (F : RO) = { + proc main () : bool = { + var b,m1,m2,hash1,hash2; + (m1,m2) <@ FEL(A,F).main(); + if (increase_counter Counter.c m1 < bound_counter) { + Counter.c <- increase_counter Counter.c m1; + hash1 <@ F.get(m1); + if (increase_counter Counter.c m2 < bound_counter) { + Counter.c <- increase_counter Counter.c m2; + hash2 <@ F.get(m2); + b <- hash1 = hash2 /\ m1 <> m2; + } else b <- false; + } + else b <- false; + return b; + } + }. + + op collision (m : ('a, 'b) fmap) = + exists m1 m2, m1 <> m2 /\ m1 \in m /\ m2 \in m /\ m.[m1] = m.[m2]. + + lemma RO_is_collision_resistant &m : + Pr [ Collision(A,RO).main() @ &m : res ] + <= ((bound * (bound - 1) + 2)%r / 2%r * mu1 sampleto witness). + proof. + have->: Pr [ Collision(A,RO).main() @ &m : res ] = + Pr [ Collision2(A,RO).main() @ &m : res ]. + + by byequiv=>//=; proc; inline*; sim. + byphoare=> //; proc. + seq 1 : (collision RO.m) + ((bound * (bound - 1))%r / 2%r * mu1 sampleto witness) 1%r + 1%r (mu1 sampleto witness) + (card (fdom RO.m) <= Bounder.bounder <= bound); first last; first last. + + auto. + + auto. + + inline*. + if; sp; last first. + - by hoare; auto; smt(mu_bounded). + case: (increase_counter Counter.c m2 < bound_counter); last first. + - rcondf 4; 1:auto; hoare; auto; smt(mu_bounded). + rcondt 4; 1: auto. + swap 4 -3. + case: (m1 = m2). + - by hoare; auto. + case: (m1 \in RO.m); case: (m2 \in RO.m). + - rcondf 3; 1: auto; rcondf 6; 1: auto; hoare; auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 _ _ _ _. + move: Hcoll; rewrite /collision negb_exists /= => /(_ m1{h}). + rewrite negb_exists /= => /(_ m2{h}). + by rewrite neq12 in_dom1 in_dom2 /=; smt(). + - rcondf 3; 1: auto; rcondt 6; 1: auto; wp; rnd (pred1 hash1). + auto=> /> &h d Hmc Hcb Hcoll _ _ neq12 in_dom1 in_dom2 _ _; split. + * smt(sampleto_fu). + by move=> _ sample _; rewrite get_set_sameE; smt(). + - rcondt 3; 1: auto; rcondf 7; 1: (by auto; smt(mem_set)). + swap 6 -5; wp=> /=; rnd (pred1 (oget RO.m.[m2])); auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 _ _; split. + * smt(sampleto_fu). + move=> _ sample _. + by rewrite get_set_sameE get_set_neqE 1:eq_sym. + rcondt 3; 1: auto; rcondt 7; 1: (by auto; smt(mem_set)). + swap 2 -1; swap 6 -4; wp=> //=; rnd (pred1 r); auto. + move=> /> &h d _ _ Hcoll _ _ neq12 in_dom1 in_dom2 sample1 _; split. + * smt(sampleto_fu). + move=> _ sample2 _. + by rewrite get_set_sameE get_set_sameE; smt(). + + by move=> />; smt(mu_bounded). + + inline*; wp; call(: card (fdom RO.m) <= Bounder.bounder <= bound); auto. + - proc; inline*; sp; if; auto; sp; if; last by auto; smt(). + auto=> /> &h d Hbc _ Hcb _ sample _; split. + * by move=> nin_dom1; rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + by move=> in_dom1; smt(). + by move=> />; rewrite fdom0 fcards0; smt(bound_gt0). + call(: true ==> collision RO.m); auto; bypr=> /> {&m} &m. + fel 1 Bounder.bounder (fun i, i%r * mu1 sampleto witness) bound + (collision RO.m) + [Bounder(RO).get: (card (fdom RO.m) <= Bounder.bounder < bound)] + (card (fdom RO.m) <= Bounder.bounder <= bound)=> //. + + rewrite -StdBigop.Bigreal.BRA.mulr_suml RField.mulrAC. + rewrite StdOrder.RealOrder.ler_wpmul2r; 1: smt(mu_bounded). + by rewrite StdBigop.Bigreal.sumidE //; smt(bound_gt0). + + inline*; auto=> />. + rewrite fdom0 fcards0; split; 2: smt(bound_gt0). + rewrite /collision negb_exists /= => a; rewrite negb_exists /= => b. + by rewrite mem_empty. + + bypr=> /> {&m} &m; pose c := Bounder.bounder{m}; move=> H0c Hcbound Hcoll Hmc _. + byphoare(: !collision RO.m /\ card (fdom RO.m) <= c ==> _)=>//=. + proc; inline*; sp; if; last first. + - by hoare; auto; smt(mu_bounded). + sp; if; last by hoare; auto; smt(mu_bounded). + case: (x \in RO.m). + - by hoare; auto; smt(mu_bounded). + rcondt 4; 1: auto; sp; wp=> /=. + conseq(:_==> r \in frng RO.m). + - move=> /> &h d c2 Hcoll2 Hb2c Hc2b _ nin_dom sample m1 m2 neq. + rewrite 2!mem_set. + case: (m1 = x{h}) => //=. + * move=> <<-; rewrite eq_sym neq /= get_set_sameE get_set_neqE//= 1:eq_sym //. + by rewrite mem_frng rngE /= => _ ->; exists m2. + case: (m2 = x{h}) => //=. + * move=> <<- _ in_dom1. + by rewrite get_set_neqE // get_set_sameE mem_frng rngE/= => <-; exists m1. + move=> neq2 neq1 in_dom1 in_dom2; rewrite get_set_neqE // get_set_neqE //. + have:= Hcoll2; rewrite negb_exists /= => /(_ m1). + rewrite negb_exists /= => /(_ m2). + by rewrite neq in_dom1 in_dom2 /= => ->. + rnd; skip=> /> &h c0 bounder _ h _. + rewrite (mu_mem (frng RO.m{h}) sampleto (mu1 sampleto witness)); 1: smt(sampleto_fu). + rewrite StdOrder.RealOrder.ler_wpmul2r //. + by rewrite le_fromint; smt(le_card_frng_fdom). + + move=> c; proc; sp; if; auto; inline*; auto; sp; if; last by auto; smt(). + auto=> /> &h d h1 _ h2 sample _. + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + move=> b c; proc; inline*; sp; if; auto; sp; if; auto; 2: smt(). + move=> /> &h h1 h2 _ h3 _ sample _. + by rewrite fdom_set fcardU fcard1; smt(fcard_ge0). + qed. + + +end section Collision. diff --git a/sha3/proof/Sponge.ec b/sha3/proof/Sponge.ec new file mode 100644 index 0000000..7ad47bb --- /dev/null +++ b/sha3/proof/Sponge.ec @@ -0,0 +1,2163 @@ +(*------------------------- Sponge Construction ------------------------*) +require import Core Int IntDiv Real List FSet SmtMap. +require import Distr DBool DList. +require import Ring StdBigop StdOrder. import IntID IntOrder. +require import Common PROM. +require (*--*) IRO BlockSponge. + +(*------------------------- Indifferentiability ------------------------*) + +clone include Indifferentiability with + type p <- block * capacity, + type f_in <- bool list * int, + type f_out <- bool list + + rename + [module] "Indif" as "Experiment" + [module] "GReal" as "RealIndif" + [module] "GIdeal" as "IdealIndif". + +(*------------------------- Ideal Functionality ------------------------*) + +clone import IRO as BIRO with + type from <- bool list, + type to <- bool, + op valid <- valid_toplevel, + op dto <- dbool. + +(*------------------------- Sponge Construction ------------------------*) + +module StdSponge (P : DPRIMITIVE) = { + proc init() : unit = {} + + proc f(bs : bool list, n : int) : bool list = { + var z <- []; + var (sa, sc) <- (b0, Capacity.c0); + var finished <- false; + var xs <- pad2blocks bs; + + (* absorption *) + while (xs <> []) { + (sa, sc) <@ P.f(sa +^ head b0 xs, sc); + xs <- behead xs; + } + (* squeezing *) + while (!finished) { + z <- z ++ ofblock sa; + if (size z < n) { + (sa, sc) <@ P.f(sa, sc); + } else { + finished <- true; + } + } + + return take n z; + } +}. + +module (Sponge : CONSTRUCTION) (P : DPRIMITIVE) : FUNCTIONALITY = { + proc init() : unit = {} + + proc f(bs : bool list, n : int) : bool list = { + var z <- []; + var (sa, sc) <- (b0, Capacity.c0); + var i <- 0; + var xs <- pad2blocks bs; + + (* absorption *) + while (xs <> []) { + (sa, sc) <@ P.f(sa +^ head b0 xs, sc); + xs <- behead xs; + } + (* squeezing *) + while (i < (n + r - 1) %/ r) { + z <- z ++ ofblock sa; + i <- i + 1; + if (i < (n + r - 1) %/ r) { + (sa, sc) <@ P.f(sa, sc); + } + } + + return take n z; + } +}. + +lemma loop_cond i n: 0 <= i => 0 <= n => r * i < n <=> i < (n + r - 1) %/ r. +proof. +move=> ge0_i; elim: i ge0_i n=> /= [|i ge0_i ih n ge0_n]. ++ smt(gt0_n). +case: (r %| n). ++ move=> ^/dvdzE n_mod_r /needed_blocks_eq_div_r <-. + by rewrite -(ltr_pmul2r r gt0_r (i + 1)) divzE n_mod_r /#. +move=> r_ndvd_n. rewrite -ltr_subr_addr -(addzC (-1)). +rewrite -divzMDr 1:#smt:(gt0_r) Ring.IntID.mulN1r. +have ->: n + r - 1 - r = (n - r) + r - 1 by smt(). +case: (0 <= n - r)=> [n_ge_r|/ltzNge n_lt_r /#]. +by rewrite -ih /#. +qed. + +equiv Sponge_is_StdSponge (P <: DPRIMITIVE): + StdSponge(P).f ~ Sponge(P).f: ={glob P, bs, n} ==> ={glob P, res}. +proof. +proc; seq 5 5: (={glob P, z, sa, sc, xs, n} /\ !finished{1} /\ i{2} = 0 /\ z{1} = []). ++ while (={glob P, xs, sa, sc}); 2:by auto. + by auto; call (: true). +case: (n{1} <= 0). ++ rcondt{1} 1=> //=; rcondf{1} 2. + + by auto; smt(size_ge0). + rcondf{1} 3; 1:by auto. + rcondf{2} 1. + + by auto=> /> &hr /needed_blocks_non_pos /#. + by auto=> /> &1 &2 _ n_le0; rewrite !take_le0. +while ( ={glob P, z, n, sa, sc} + /\ (finished{1} <=> n{1} <= size z{1}) + /\ size z{1} = r * i{2} + /\ 0 < n{1} + /\ 0 <= i{2}). ++ sp; if=> />. + + move=> &2 i z; rewrite size_cat size_block=> -> gt0_n ge0_i /ltzNge gt_ri_n gt_i_nbl. + by rewrite -(mulzDr r i 1) loop_cond /#. + + call (: true); auto=> /> &2 i z; rewrite size_cat size_block=> -> gt0_n ge0_i /ltzNge gt_ri_n gt_i_nbl. + move=> ^ + /ltzNge -> /=; rewrite mulzDr /=. + by rewrite -(mulzDr r i 1) loop_cond /#. + + auto=> /> &2 i z; rewrite size_cat size_block=> -> gt0_n ge0_i /ltzNge gt_ri_n gt_i_nbl. + move=> ^ + /lezNgt -> /=; rewrite mulzDr /=. + by rewrite -(mulzDr r i 1) loop_cond /#. +by auto=> /> &1 &2 _ /ltrNge gt0_n; smt(gt0_r). +qed. + +(*------------- Simulator and Distinguisher Constructions --------------*) + +module LowerFun (F : DFUNCTIONALITY) : BlockSponge.DFUNCTIONALITY = { + proc f(xs : block list, n : int) = { + var cs, ds : bool list; + var obs : bool list option; + var ys : block list <- []; + + obs <- unpad_blocks xs; + if (obs <> None) { + cs <@ F.f(oget obs, n * r); (* size cs = n * r *) + ys <- bits2blocks cs; (* size ys = n *) + } + return ys; + } +}. + +module RaiseFun (F : BlockSponge.DFUNCTIONALITY) : DFUNCTIONALITY = { + proc f(bs : bool list, n : int) = { + var xs; + + xs <@ F.f(pad2blocks bs, (n + r - 1) %/ r); + return take n (blocks2bits xs); + } +}. + +module LowerDist (D : DISTINGUISHER, F : BlockSponge.DFUNCTIONALITY) = + D(RaiseFun(F)). + +module RaiseSim (S : BlockSponge.SIMULATOR, F : DFUNCTIONALITY) = + S(LowerFun(F)). + +(* Our main result will be: + + lemma conclusion + (BlockSim <: BlockSponge.SIMULATOR{-IRO, -BlockSponge.BIRO.IRO}) + (Dist <: DISTINGUISHER{-Perm, -BlockSim, -IRO, -BlockSponge.BIRO.IRO}) + &m : + `|Pr[RealIndif(Sponge, Perm, Dist).main() @ &m : res] - + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res]| = + `|Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, LowerDist(Dist)).main() @ &m : res] - + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main() @ &m : res]| +*) + +(*------------------------------- Proof --------------------------------*) + +(* Proving the Real side + + Pr[RealIndif(Sponge, Perm, Dist).main() @ &m : res] = + Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, LowerDist(Dist)).main() @ &m : res] + + is easy (see lemma RealIndif_Sponge_BlockSponge) + + And we split the proof of the Ideal side (IdealIndif_IRO_BlockIRO) + + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res] = + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main () @ &m : res]. + + into three steps, involving Hybrid IROs, which, in addition to + an init procedure, have the procedure + + (* hashing block lists, giving n bits *) + proc f(x : block list, n : int) : bool list + + We have lazy (HybridIROLazy) and eager (HybridIROEager) Hybrid + IROs, both of which work with a finite map from block list * int to + bool. In both versions, as in BlockSponge.BIRO.IRO, f returns [] if + x isn't a valid block list. + + In the lazy version, f consults/randomly updates just those + elements of the map's domain needed to produce the needed bits. But + the eager version goes further, consulting/randomly updating enough + extra domain elements so that a multiple of r domain elements are + consulted/randomly updated (those extra bits are discarded) + + We have a parameterized module RaiseHybridIRO for turning a Hybrid + IRO into a FUNCTIONALITY in the obvious way, and we have a + parameterized module LowerHybridIRO for turning a Hybrid IRO into a + DFUNCTIONALITY in the obivous way. We split the proof of the Ideal + side into three steps: + + Step 1: + + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res] = + Pr[Experiment + (RaiseHybridIRO(HybridIROLazy), + BlockSim(LowerHybridIRO(HybridIROLazy)), + Dist).main() @ &m : res] + + This step is proved using a lazy invariant relating the + maps of the bit-based IRO and HybridIROLazy + + Step 2: + + Pr[Experiment + (RaiseHybridIRO(HybridIROLazy), + BlockSim(LowerHybridIRO(HybridIROLazy)), + Dist).main() @ &m : res] = + Pr[Experiment + (RaiseHybridIRO(HybridIROEager), + BlockSim(LowerHybridIRO(HybridIROEager)), + Dist).main() @ &m : res] + + This step is proved using the eager sampling lemma provided by + PROM. + + Step 3: + + Pr[Experiment + (RaiseHybridIRO(HybridIROEager), + BlockSim(LowerHybridIRO(HybridIROEager)), + Dist).main() @ &m : res] = + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main () @ &m : res] + + This step is proved using an invariant relating the maps of + HybridIROEager and the block-based IRO. Its proof is the most + involved, and uses the Program abstract theory of DList to show the + equivalence of randomly choosing a block and forming a block out + of r randomly chosen bits *) + +(*------------------- abstract theory of Hybrid IROs -------------------*) + +abstract theory HybridIRO. + +module type HYBRID_IRO = { + (* initialization *) + proc init() : unit + + (* hashing block lists, giving n bits *) + proc f(x : block list, n : int) : bool list +}. + +(* distinguisher for Hybrid IROs *) + +module type HYBRID_IRO_DIST(HI : HYBRID_IRO) = { + proc distinguish() : bool {HI.f} +}. + +(* experiments for Hybrid IROs *) + +module HybridIROExper(HI : HYBRID_IRO, D : HYBRID_IRO_DIST) = { + proc main() : bool = { + var b : bool; + HI.init(); + b <@ D(HI).distinguish(); + return b; + } +}. + +(* lazy implementation of Hybrid IROs *) + +module HybridIROLazy : HYBRID_IRO = { + var mp : (block list * int, bool) fmap + + proc init() : unit = { + mp <- empty; + } + + proc fill_in(xs : block list, i : int) = { + var r; + + if (! dom mp (xs, i)) { + r <$ dbool; + mp.[(xs, i)] <- r; + } + return oget mp.[(xs, i)]; + } + + proc f(xs : block list, n : int) = { + var b, bs; + var i <- 0; + + bs <- []; + if (valid_block xs) { + while (i < n) { + b <@ fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + } + return bs; + } +}. + +(* eager implementation of Hybrid IROs *) + +module HybridIROEager : HYBRID_IRO = { + (* same as lazy implementation, except for indicated part *) + var mp : (block list * int, bool) fmap + + proc init() : unit = { + mp <- empty; + } + + proc fill_in(xs : block list, i : int) = { + var r; + + if (! dom mp (xs, i)) { + r <$ dbool; + mp.[(xs, i)] <- r; + } + return oget mp.[(xs, i)]; + } + + proc f(xs : block list, n : int) = { + var b, bs; + var m <- ((n + r - 1) %/ r) * r; (* eager part *) + var i <- 0; + + bs <- []; + if (valid_block xs) { + while (i < n) { + b <@ fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + while (i < m) { (* eager part *) + fill_in(xs, i); + i <- i + 1; + } + } + return bs; + } +}. + +(* we are going to use PROM.GenEager to prove: + +lemma HybridIROExper_Lazy_Eager + (D <: HYBRID_IRO_DIST{-HybridIROEager, -HybridIROLazy}) &m : + Pr[HybridIROExper(HybridIROLazy, D).main() @ &m : res] = + Pr[HybridIROExper(HybridIROEager, D).main() @ &m : res]. +*) + +section. + +declare module D <: HYBRID_IRO_DIST {-HybridIROEager, -HybridIROLazy}. + +local clone import PROM.FullRO as ERO with + type in_t <- block list * int, + type out_t <- bool, + op dout _ <- dbool, + type d_in_t <- unit, + type d_out_t <- bool. + +local module EROExper(O : ERO.RO, D : ERO.RO_Distinguisher) = { + proc main() : bool = { + var b : bool; + O.init(); + b <@ D(O).distinguish(); + return b; + } +}. + +local lemma LRO_RO (D <: ERO.RO_Distinguisher{-ERO.RO, -ERO.FRO}) &m : + Pr[EROExper(LRO, D).main() @ &m : res] = + Pr[EROExper(ERO.RO, D).main() @ &m : res]. +proof. +byequiv=> //; proc. +seq 1 1 : (={glob D, ERO.RO.m}); first sim. +by symmetry; call (FullEager.RO_LRO_D D _); auto; rewrite dbool_ll. +qed. + +(* make a Hybrid IRO out of a random oracle *) + +local module HIRO(RO : ERO.RO) : HYBRID_IRO = { + proc init() : unit = { + RO.init(); + } + + proc f(xs, n) = { + var b, bs; + var m <- ((n + r - 1) %/ r) * r; + var i <- 0; + + bs <- []; + if (valid_block xs) { + while (i < n) { + b <@ RO.get(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + while (i < m) { + RO.sample(xs, i); + i <- i + 1; + } + } + return bs; + } +}. + +local lemma HybridIROLazy_HIRO_LRO_init : + equiv[HybridIROLazy.init ~ HIRO(LRO).init : + true ==> HybridIROLazy.mp{1} = ERO.RO.m{2}]. +proof. proc; inline*; auto. qed. + +local lemma HybridIROLazy_fill_in_LRO_get : + equiv[HybridIROLazy.fill_in ~ LRO.get : + (xs, i){1} = x{2} /\ HybridIROLazy.mp{1} = ERO.RO.m{2} ==> + ={res} /\ HybridIROLazy.mp{1} = ERO.RO.m{2}]. +proof. +proc=> /=. +case ((dom HybridIROLazy.mp{1}) (xs{1}, i{1})). +rcondf{1} 1; first auto. rcondf{2} 2; first auto. +rnd{2}; auto; progress; apply dbool_ll. +rcondt{1} 1; first auto. rcondt{2} 2; first auto. +wp; rnd; auto. +qed. + +local lemma HybridIROLazy_HIRO_LRO_f : + equiv[HybridIROLazy.f ~ HIRO(LRO).f : + ={xs, n} /\ HybridIROLazy.mp{1} = ERO.RO.m{2} ==> + ={res} /\ HybridIROLazy.mp{1} = ERO.RO.m{2}]. +proof. +proc; inline LRO.sample; sp=> /=. +if=> //. +while{2} (true) (m{2} - i{2}). +progress; auto; progress; smt(). +while (={xs, n, i, bs} /\ HybridIROLazy.mp{1} = ERO.RO.m{2}). +wp; call HybridIROLazy_fill_in_LRO_get; auto. +auto; progress; smt(). +qed. + +local lemma HIRO_RO_HybridIROEager_init : + equiv[HIRO(ERO.RO).init ~ HybridIROEager.init : + true ==> ={res} /\ ERO.RO.m{1} = HybridIROEager.mp{2}]. +proof. proc; inline*; auto. qed. + +local lemma RO_get_HybridIROEager_fill_in : + equiv[ERO.RO.get ~ HybridIROEager.fill_in : + x{1} = (xs, i){2} /\ ERO.RO.m{1} = HybridIROEager.mp{2} ==> + ={res} /\ ERO.RO.m{1} = HybridIROEager.mp{2}]. +proof. +proc=> /=. +case (dom HybridIROEager.mp{2} (xs{2}, i{2})). +rcondf{1} 2; first auto. rcondf{2} 1; first auto. +rnd{1}; auto; progress; apply dbool_ll. +rcondt{1} 2; first auto. rcondt{2} 1; first auto. +wp; rnd; auto. +qed. + +local lemma RO_sample_HybridIROEager_fill_in : + equiv[ERO.RO.sample ~ HybridIROEager.fill_in : + x{1} = (xs, i){2} /\ ERO.RO.m{1} = HybridIROEager.mp{2} ==> + ERO.RO.m{1} = HybridIROEager.mp{2}]. +proof. +proc=> /=; inline ERO.RO.get; sp. +case (dom HybridIROEager.mp{2} (xs{2}, i{2})). +rcondf{1} 2; first auto. rcondf{2} 1; first auto. +rnd{1}; auto; progress; apply dbool_ll. +rcondt{1} 2; first auto. rcondt{2} 1; first auto. +wp; rnd; auto. +qed. + +local lemma HIRO_RO_HybridIROEager_f : + equiv[HIRO(ERO.RO).f ~ HybridIROEager.f : + ={xs, n} /\ ERO.RO.m{1} = HybridIROEager.mp{2} ==> + ={res} /\ ERO.RO.m{1} = HybridIROEager.mp{2}]. +proof. +proc; first sp=> /=. +if=> //. +while (={i, m, xs} /\ ERO.RO.m{1} = HybridIROEager.mp{2}). +wp; call RO_sample_HybridIROEager_fill_in; auto. +while (={i, n, xs, bs} /\ ERO.RO.m{1} = HybridIROEager.mp{2}). +wp; call RO_get_HybridIROEager_fill_in; auto. +auto. +qed. + +(* make distinguisher for random oracles out of HIRO and D *) + +local module RODist(RO : ERO.RO) = { + proc distinguish() : bool = { + var b : bool; + b <@ D(HIRO(RO)).distinguish(); + return b; + } +}. + +local lemma Exper_HybridIROLazy_LRO &m : + Pr[HybridIROExper(HybridIROLazy, D).main() @ &m : res] = + Pr[EROExper(LRO, RODist).main() @ &m : res]. +proof. +byequiv=> //; proc; inline*; wp. +seq 1 1 : (={glob D} /\ HybridIROLazy.mp{1} = ERO.RO.m{2}); first auto. +call (_ : HybridIROLazy.mp{1} = ERO.RO.m{2}). +conseq HybridIROLazy_HIRO_LRO_f. +auto. +qed. + +local lemma Exper_RO_HybridIROEager &m : + Pr[EROExper(ERO.RO, RODist).main() @ &m : res] = + Pr[HybridIROExper(HybridIROEager, D).main() @ &m : res]. +proof. +byequiv=> //; proc; inline*; wp. +seq 1 1 : (={glob D} /\ ERO.RO.m{1} = HybridIROEager.mp{2}); first auto. +call (_ : ERO.RO.m{1} = HybridIROEager.mp{2}). +conseq HIRO_RO_HybridIROEager_f. +auto. +qed. + +lemma HybridIROExper_Lazy_Eager' &m : + Pr[HybridIROExper(HybridIROLazy, D).main() @ &m : res] = + Pr[HybridIROExper(HybridIROEager, D).main() @ &m : res]. +proof. +by rewrite (Exper_HybridIROLazy_LRO &m) + (LRO_RO RODist &m) + (Exper_RO_HybridIROEager &m). +qed. + +end section. + +lemma HybridIROExper_Lazy_Eager + (D <: HYBRID_IRO_DIST{-HybridIROEager, -HybridIROLazy}) &m : + Pr[HybridIROExper(HybridIROLazy, D).main() @ &m : res] = + Pr[HybridIROExper(HybridIROEager, D).main() @ &m : res]. +proof. by apply (HybridIROExper_Lazy_Eager' D &m). qed. + +(* turn a Hybrid IRO implementation (lazy or eager) into top-level + ideal functionality *) + +module RaiseHybridIRO (HI : HYBRID_IRO) : FUNCTIONALITY = { + proc init() = { + HI.init(); + } + + proc f(bs : bool list, n : int) = { + var cs; + cs <@ HI.f(pad2blocks bs, n); + return cs; + } +}. + +(* turn a Hybrid IRO implementation (lazy or eager) into lower-level + ideal distinguisher functionality *) + +module LowerHybridIRO (HI : HYBRID_IRO) : BlockSponge.DFUNCTIONALITY = { + proc f(xs : block list, n : int) = { + var bs, ys; + bs <@ HI.f(xs, n * r); + ys <- bits2blocks bs; + return ys; + } +}. + +(* invariant relating maps of BIRO.IRO and HybridIROLazy *) + +pred lazy_invar + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap) = + (forall (bs : bool list, n : int), + dom mp1 (bs, n) <=> dom mp2 (pad2blocks bs, n)) /\ + (forall (xs : block list, n), + dom mp2 (xs, n) => valid_block xs) /\ + (forall (bs : bool list, n : int), + dom mp1 (bs, n) => + oget mp1.[(bs, n)] = oget mp2.[(pad2blocks bs, n)]). + +lemma lazy_invar0 : lazy_invar empty empty. +proof. +split; first smt(mem_empty). +split; first smt(mem_empty). +smt(mem_empty). +qed. + +lemma lazy_invar_mem_pad2blocks_l2r + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs : bool list, i : int) : + lazy_invar mp1 mp2 => dom mp1 (bs, i) => + dom mp2 (pad2blocks bs, i). +proof. smt(). qed. + +lemma lazy_invar_mem_pad2blocks_r2l + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs : bool list, i : int) : + lazy_invar mp1 mp2 => dom mp2 (pad2blocks bs, i) => + dom mp1 (bs, i). +proof. smt(). qed. + +lemma lazy_invar_vb + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, n : int) : + lazy_invar mp1 mp2 => dom mp2 (xs, n) => + valid_block xs. +proof. smt(). qed. + +lemma lazy_invar_lookup_eq + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs : bool list, n : int) : + lazy_invar mp1 mp2 => dom mp1 (bs, n) => + oget mp1.[(bs, n)] = oget mp2.[(pad2blocks bs, n)]. +proof. smt(). qed. + +lemma lazy_invar_upd_mem_dom_iff + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs cs : bool list, n m : int, b : bool) : + lazy_invar mp1 mp2 => + dom mp1.[(bs, n) <- b] (cs, m) <=> + dom mp2.[(pad2blocks bs, n) <- b] (pad2blocks cs, m). +proof. +move=> li; split=> [mem_upd_mp1 | mem_upd_mp2]. +rewrite mem_set; rewrite mem_set in mem_upd_mp1. +case: ((cs, m) = (bs, n))=> [cs_m_eq_bs_n | cs_m_neq_bs_n]. +right; by elim cs_m_eq_bs_n=> -> ->. +left; smt(). +rewrite mem_set; rewrite mem_set in mem_upd_mp2. +case: ((cs, m) = (bs, n))=> [// | cs_m_neq_bs_n]. +elim mem_upd_mp2=> [/# | [p2b_cs_p2b_bs eq_mn]]. +have /# : cs = bs by apply pad2blocks_inj. +qed. + +lemma lazy_invar_upd2_vb + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs : bool list, xs : block list, n m : int, b : bool) : + lazy_invar mp1 mp2 => + dom mp2.[(pad2blocks bs, n) <- b] (xs, m) => + valid_block xs. +proof. +move=> li mem_upd_mp2. +rewrite mem_set in mem_upd_mp2. +elim mem_upd_mp2=> [/# | [-> _]]. +apply valid_pad2blocks. +qed. + +lemma lazy_invar_upd_lu_eq + (mp1 : (bool list * int, bool) fmap, + mp2 : (block list * int, bool) fmap, + bs cs : bool list, n m : int, b : bool) : + lazy_invar mp1 mp2 => + dom mp1.[(bs, n) <- b] (cs, m) => + oget mp1.[(bs, n) <- b].[(cs, m)] = + oget mp2.[(pad2blocks bs, n) <- b].[(pad2blocks cs, m)]. +proof. +move=> li mem_upd_mp1. +case: ((cs, m) = (bs, n))=> [[-> ->] | cs_m_neq_bs_n]. ++ by rewrite !get_set_sameE. +rewrite mem_set in mem_upd_mp1. +elim mem_upd_mp1=> [mem_mp1 | [-> ->]]. ++ case: ((pad2blocks bs, n) = (pad2blocks cs, m))=> + [[p2b_bs_p2b_cs ->>] | p2b_bs_n_neq_p2b_cs_m]. + + move: (pad2blocks_inj _ _ p2b_bs_p2b_cs)=> ->>. + by move: cs_m_neq_bs_n=> //=. + rewrite !get_set_neqE 1:// 1:eq_sym //. + by move: li=> [] _ [] _ /(_ _ _ mem_mp1). +by rewrite !get_set_sameE. +qed. + +lemma LowerFun_IRO_HybridIROLazy_f : + equiv[LowerFun(IRO).f ~ LowerHybridIRO(HybridIROLazy).f : + ={xs, n} /\ lazy_invar IRO.mp{1} HybridIROLazy.mp{2} ==> + ={res} /\ lazy_invar IRO.mp{1} HybridIROLazy.mp{2}]. +proof. +proc=> /=; inline HybridIROLazy.f. +seq 0 1 : + (={n} /\ xs{1} = xs0{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2}); first auto. +case (valid_block xs{1}). +rcondt{1} 3; first auto. rcondt{2} 4; first auto. +inline*. rcondt{1} 7; first auto. +seq 6 3 : + (={i, n0} /\ bs{1} = bs0{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2} /\ + pad2blocks x{1} = xs0{2}). +auto; progress. + have {2}<- := unpadBlocksK xs0{2}; first + by rewrite (some_oget (unpad_blocks xs0{2})). +wp. +while + (={i, n0} /\ bs{1} = bs0{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2} /\ + pad2blocks x{1} = xs0{2}). +sp; auto. +if. +progress; + [by apply (lazy_invar_mem_pad2blocks_l2r IRO.mp{1} + HybridIROLazy.mp{2} x{1} i{2}) | + by apply (lazy_invar_mem_pad2blocks_r2l IRO.mp{1} + HybridIROLazy.mp{2} x{1} i{2})]. +wp; rnd; auto=> |> &1 &2 inv i_lt_n0 xi_notin_m r _. +rewrite !get_set_sameE //=; split=> [bs n|]. ++ exact/(lazy_invar_upd_mem_dom_iff _ _ _ _ _ _ _ inv). +split=> [xs n|bs n]. ++ by move=>/(lazy_invar_upd2_vb _ _ _ _ _ _ _ inv). +by move=>/(lazy_invar_upd_lu_eq _ _ _ _ _ _ _ inv). +auto; progress [-delta]. +by rewrite (lazy_invar_lookup_eq IRO.mp{1} HybridIROLazy.mp{2} x{1} i{2}). +auto. +rcondf{1} 3; first auto. rcondf{2} 4; first auto. +auto; progress; by rewrite bits2blocks_nil. +qed. + +lemma IRO_RaiseHybridIRO_HybridIROLazy_f : + equiv[IRO.f ~ RaiseHybridIRO(HybridIROLazy).f : + ={n} /\ x{1} = bs{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2} ==> + ={res} /\ lazy_invar IRO.mp{1} HybridIROLazy.mp{2}]. +proof. +proc=> /=; inline*. +rcondt{1} 3; first auto. +rcondt{2} 5; first auto; progress; apply valid_pad2blocks. +seq 2 4 : + (={i, n} /\ n{1} = n0{2} /\ xs{2} = pad2blocks x{1} /\ bs{1} = bs0{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2}); first auto. +wp. +while + (={i, n} /\ n{1} = n0{2} /\ xs{2} = pad2blocks x{1} /\ bs{1} = bs0{2} /\ + lazy_invar IRO.mp{1} HybridIROLazy.mp{2}). +wp; sp. +if. +progress; + [by apply (lazy_invar_mem_pad2blocks_l2r IRO.mp{1} + HybridIROLazy.mp{2} x{1} i{2}) | + by apply (lazy_invar_mem_pad2blocks_r2l IRO.mp{1} + HybridIROLazy.mp{2} x{1} i{2})]. +wp; rnd; auto=> |> &1 &2 inv i_lt_n0 xi_notin_m r _. +rewrite !get_set_sameE=> //=; split=> [bs n|]. ++ exact/(lazy_invar_upd_mem_dom_iff _ _ _ _ _ _ _ inv). +split=> [xs n|bs n]. ++ by move=>/(lazy_invar_upd2_vb _ _ _ _ _ _ _ inv). +by move=>/(lazy_invar_upd_lu_eq _ _ _ _ _ _ _ inv). +auto; progress [-delta]; + by rewrite (lazy_invar_lookup_eq IRO.mp{1} HybridIROLazy.mp{2} x{1} i{2}). +auto. +qed. + +(* invariant relating maps of BlockSponge.BIRO.IRO and HybridIROEager *) + +pred eager_invar + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap) = + (forall (xs : block list, i : int), + dom mp1 (xs, i) => + 0 <= i /\ + (forall (j : int), i * r <= j < (i + 1) * r => + mp2.[(xs, j)] = + Some(nth false (ofblock (oget mp1.[(xs, i)])) (j - i * r)))) /\ + (forall (xs : block list, j : int), + dom mp2 (xs, j) => dom mp1 (xs, j %/ r)). + +pred block_bits_all_in_dom + (xs : block list, i : int, mp : (block list * int, bool) fmap) = + forall (j : int), i <= j < i + r => dom mp (xs, j). + +pred block_bits_all_out_dom + (xs : block list, i : int, mp : (block list * int, bool) fmap) = + forall (j : int), i <= j < i + r => ! dom mp (xs, j). + +pred block_bits_dom_all_in_or_out + (xs : block list, i : int, mp : (block list * int, bool) fmap) = + block_bits_all_in_dom xs i mp \/ block_bits_all_out_dom xs i mp. + +lemma eager_inv_mem_mp1_ge0 + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, i : int) : + eager_invar mp1 mp2 => dom mp1 (xs, i) => 0 <= i. +proof. move=> [ei1 ei2] mem_mp1_i; smt(). qed. + +lemma eager_inv_mem_mp2_ge0 + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, j : int) : + eager_invar mp1 mp2 => dom mp2 (xs, j) => 0 <= j. +proof. +move=> [ei1 ei2] mem_mp2_j. +have mem_mp1_j_div_r : dom mp1 (xs, j %/ r) by smt(). +have ge0_j_div_r : 0 <= j %/ r by smt(). +smt(divz_ge0 gt0_r). +qed. + +lemma eager_invar0 : eager_invar empty empty. +proof. split; smt(mem_empty). qed. + +lemma eager_inv_imp_block_bits_dom + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, i : int) : + 0 <= i => r %| i => eager_invar mp1 mp2 => + block_bits_dom_all_in_or_out xs i mp2. +proof. +move=> ge0_i r_dvd_i [ei1 ei2]. +case: (dom mp1 (xs, i %/ r))=> [mem_mp1 | not_mem_mp1]. +have ei1_xs_i_div_r := ei1 xs (i %/ r). +have [_ mp2_eq_block_bits] := ei1_xs_i_div_r mem_mp1. +left=> j j_rng. +have mp2_eq_block_bits_j := mp2_eq_block_bits j _. + by rewrite divzK // mulzDl /= divzK. +rewrite domE /#. +right=> j j_rng. +case: (dom mp2 (xs, j))=> // mem_mp2 /=. +have mem_mp1 := ei2 xs j mem_mp2. +have [k] [k_ran j_eq_i_plus_k] : exists k, 0 <= k < r /\ j = i + k + by exists (j - i); smt(). +have /# : (i + k) %/r = i %/ r + by rewrite divzDl // (divz_small k r) 1:ger0_norm 1:ge0_r. +qed. + +lemma eager_inv_mem_dom1 + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, i : int) : + eager_invar mp1 mp2 => dom mp1 (xs, i) => + block_bits_all_in_dom xs (i * r) mp2. +proof. +move=> [ei1 _] mem j j_ran. +have [ge0_i eq_mp2_block_i] := ei1 xs i mem. +rewrite domE. +have /# := eq_mp2_block_i j _; smt(). +qed. + +lemma eager_inv_not_mem_dom1 + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + xs : block list, i : int) : + eager_invar mp1 mp2 => 0 <= i => ! dom mp1 (xs, i) => + block_bits_all_out_dom xs (i * r) mp2. +proof. +move=> [_ ei2] ge0_i not_mem_mp1_i j j_ran. +case (dom mp2 (xs, j))=> // mem_mp2_j. +have mem_mp1_j_div_r := ei2 xs j mem_mp2_j. +have /# : j %/ r = i. +have [k] [k_ran ->] : exists k, 0 <= k < r /\ j = i * r + k + by exists (j - i * r); smt(). +by rewrite divzDl 1:dvdz_mull 1:dvdzz (divz_small k r) + 1:ger0_norm 1:ge0_r //= mulzK 1:gtr_eqF 1:gt0_r. +qed. + +lemma block_bits_dom_first_in_imp_all_in + (xs : block list, i : int, mp : (block list * int, bool) fmap) : + block_bits_dom_all_in_or_out xs i mp => dom mp (xs, i) => + block_bits_all_in_dom xs i mp. +proof. smt(). qed. + +lemma block_bits_dom_first_out_imp_all_out + (xs : block list, i : int, mp : (block list * int, bool) fmap) : + block_bits_dom_all_in_or_out xs i mp => ! dom mp (xs, i) => + block_bits_all_out_dom xs i mp. +proof. smt(). qed. + +lemma Lower_HybridIROEager_f : + equiv[LowerHybridIRO(HybridIROEager).f ~ HybridIROEager.f : + ={xs, HybridIROEager.mp} /\ n{1} * r = n{2} ==> + res{1} = bits2blocks res{2} /\ ={HybridIROEager.mp}]. +proof. +proc=> /=; inline*. +seq 5 3 : + (={i, HybridIROEager.mp} /\ xs0{1} = xs{2} /\ + bs0{1} = bs{2} /\ n0{1} = n{2} /\ m{1} = n0{1} /\ m{2} = n{2}). +auto; progress; first 2 by rewrite needed_blocks_prod_r. +if=> //; wp. +while + (={i, HybridIROEager.mp} /\ xs0{1} = xs{2} /\ + bs0{1} = bs{2} /\ n0{1} = n{2} /\ m{1} = n0{1} /\ + m{2} = n{2}). +sp; wp; if=> //; wp; rnd; auto. +while + (={i, HybridIROEager.mp} /\ xs0{1} = xs{2} /\ + bs0{1} = bs{2} /\ n0{1} = n{2} /\ m{1} = n0{1} /\ + m{2} = n{2})=> //. +sp; wp; if=> //; wp; rnd; auto. +auto. +qed. + +(* module needed for applying transitivity tactic in connection + with HybridIROEager *) + +module HybridIROEagerTrans = { + (* getting next block of bits; assuming m = i + r and size bs = i *) + + proc next_block(xs, i, m : int, bs) = { + var b; + + while (i < m) { + b <@ HybridIROEager.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + return (bs, i); + } + + (* version of next_block split into cases; assuming m = i + r, + size bs = i and block_bits_dom_all_in_or_out xs i HybridIROEager.mp *) + + proc next_block_split(xs, i, m : int, bs) = { + var b, j, cs; + + if (dom HybridIROEager.mp (xs, i)) { + while (i < m) { + b <- oget HybridIROEager.mp.[(xs, i)]; + bs <- rcons bs b; + i <- i + 1; + } + } else { + j <- 0; cs <- []; + while (j < r) { + b <$ dbool; + cs <- rcons cs b; + j <- j + 1; + } + bs <- bs ++ cs; + while (i < m) { + HybridIROEager.mp.[(xs, i)] <- nth true bs i; + i <- i + 1; + } + } + return (bs, i); + } + + (* loop getting n * r bits of hash *) + + proc loop(n : int, xs : block list) : int * bool list = { + var b : bool; var i <- 0; var bs <- []; + while (i < n * r) { + b <@ HybridIROEager.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + return (i, bs); + } +}. + +(* predicate saying two (block list * int, bool) maps are the same + except (perhaps) on a range of bits for a single block list *) + +pred eager_eq_except + (xs : block list, i j : int, + mp1 mp2 : (block list * int, bool) fmap) = + forall (ys : block list, k : int), + ys <> xs \/ k < i \/ j <= k => mp1.[(ys, k)] = mp2.[(ys, k)]. + +lemma eager_eq_except_mem_iff + (xs ys : block list, i j k : int, + mp1 mp2 : (block list * int, bool) fmap) : + eager_eq_except xs i j mp1 mp2 => + ys <> xs \/ k < i \/ j <= k => + dom mp1 (ys, k) <=> dom mp2 (ys, k). +proof. smt(domE). qed. + +lemma eager_eq_except_upd1_eq_in + (xs : block list, i j k : int, y : bool, + mp1 mp2 : (block list * int, bool) fmap) : + eager_eq_except xs i j mp1 mp2 => i <= k => k < j => + eager_eq_except xs i j mp1.[(xs, k) <- y] mp2. +proof. +move=> eee le_ik lt_kj ys l disj. +have ne : (xs, k) <> (ys, l) by smt(). +smt(get_setE). +qed. + +lemma eager_eq_except_upd2_eq_in + (xs : block list, i j k : int, y : bool, + mp1 mp2 : (block list * int, bool) fmap) : + eager_eq_except xs i j mp1 mp2 => i <= k => k < j => + eager_eq_except xs i j mp1 mp2.[(xs, k) <- y]. +proof. +move=> eee le_ik lt_kj ys l disj. +have ne : (xs, k) <> (ys, l) by smt(). +smt(get_setE). +qed. + +lemma eager_eq_except_maps_eq + (xs : block list, i j : int, y : bool, + mp1 mp2 : (block list * int, bool) fmap) : + i <= j => eager_eq_except xs i j mp1 mp2 => + (forall (k : int), + i <= k < j => mp1.[(xs, k)] = mp2.[(xs, k)]) => + mp1 = mp2. +proof. +move=> lt_ij eee ran_k. +apply fmap_eqP=> p. +have [ys k] -> /# : exists ys k, p = (ys, k) + by exists p.`1 p.`2; smt(). +qed. + +lemma eager_invar_eq_except_upd1 + (mp1 : (block list * int, block) fmap, + mp2 : (block list * int, bool) fmap, + mp2' : (block list * int, bool) fmap, + xs : block list, i : int, y : block) : + 0 <= i => eager_invar mp1 mp2 => + eager_eq_except xs (i * r) ((i + 1) * r) mp2 mp2' => + (forall (j : int), + i * r <= j < (i + 1) * r => + mp2'.[(xs, j)] = Some (nth false (ofblock y) (j - i * r))) => + eager_invar mp1.[(xs, i) <- y] mp2'. +proof. +move=> ge0_i [ei1 ei2] ee mp2'_ran_eq. +split=> [ys k mem_mp1_upd_xs_i_y_ys_k | ys k mem_dom_mp2'_ys_k]. +case: (xs = ys)=> [eq_xs_ys | ne_xs_ys]. +case: (k = i)=> [eq_k_i | ne_k_i]. +split; first smt(). +move=> j j_ran. +by rewrite -eq_xs_ys eq_k_i get_set_sameE mp2'_ran_eq -eq_k_i. +rewrite domE in mem_mp1_upd_xs_i_y_ys_k. +have xs_i_ne_ys_k : (xs, i) <> (ys, k) by smt(). +have mem_mp1_ys_k : dom mp1 (ys, k) by smt(get_setE). +split; first smt(eager_inv_mem_mp2_ge0). +move=> j j_ran; rewrite get_setE. +have -> /= : (ys, k) <> (xs, i) by smt(). +have [_ ei1_ys_k_snd] := ei1 ys k mem_mp1_ys_k. +have <- : + mp2.[(ys, j)] = + Some (nth false (ofblock (oget mp1.[(ys, k)])) (j - k * r)) + by rewrite ei1_ys_k_snd. +have /# : j < i * r \/ (i + 1) * r <= j. + have [lt_ki | lt_ik] : k < i \/ i < k by smt(). + left. + have le_k_add1_i : k + 1 <= i + by rewrite addzC lez_add1r. + by rewrite (ltr_le_trans ((k + 1) * r)) 1:/# ler_pmul2r 1:gt0_r. + right. + have le_i_add1_k : i + 1 <= k + by rewrite addzC lez_add1r. + rewrite (lez_trans (k * r)) 1:ler_pmul2r 1:gt0_r // /#. +rewrite domE in mem_mp1_upd_xs_i_y_ys_k. +have xs_i_ne_ys_k : (xs, i) <> (ys, k) by smt(). +have mem_mp1_ys_k : dom mp1 (ys, k) by smt(get_setE). +split; first smt(eager_inv_mem_mp2_ge0). +move=> j j_ran; rewrite get_setE. +have -> /= : (ys, k) <> (xs, i) by smt(). +have [_ ei1_ys_k_snd] := ei1 ys k mem_mp1_ys_k. +have <- /# : + mp2.[(ys, j)] = + Some (nth false (ofblock (oget mp1.[(ys, k)])) (j - k * r)) + by rewrite ei1_ys_k_snd. +rewrite domE. +case: (xs = ys)=> [-> | ne_xs_ys]. +case: (k < i * r)=> [lt_k_i_tim_r | not_lt_k_i_tim_r]. +smt(get_setE eager_eq_except_mem_iff). +case: ((i + 1) * r <= k)=> [i_add1_tim_r_le_k | not_i_add1_tim_r_le_k]. +smt(get_setE eager_eq_except_mem_iff). +have le_i_tim_r_k : i * r <= k by smt(). +have lt_k_i_add1_tim_r : k < (i + 1) * r by smt(). +have -> // : i = k %/ r. + apply eqz_leq; split. + by rewrite lez_divRL 1:gt0_r. + by rewrite -ltzS ltz_divLR 1:gt0_r. +smt(get_setE eager_eq_except_mem_iff). +smt(get_setE eager_eq_except_mem_iff). +qed. + +lemma HybridIROEagerTrans_next_block_split : + equiv + [HybridIROEagerTrans.next_block ~ HybridIROEagerTrans.next_block_split : + ={i, m, xs, bs, HybridIROEager.mp} /\ m{1} = i{1} + r /\ + size bs{1} = i{1} /\ + block_bits_dom_all_in_or_out xs{1} i{1} HybridIROEager.mp{1} ==> + ={res, HybridIROEager.mp}]. +proof. +proc=> /=. +case (dom HybridIROEager.mp{2} (xs{2}, i{2})). +(* dom HybridIROEager.mp{2} (xs{2}, i{2}) *) +rcondt{2} 1; first auto. +conseq + (_ : + ={i, m, xs, bs, HybridIROEager.mp} /\ i{1} <= m{1} /\ + (forall (j : int), + i{1} <= j < m{1} => + dom HybridIROEager.mp{1} (xs{1}, j)) ==> + _). +progress; smt(gt0_r). +while + (={i, m, xs, bs, HybridIROEager.mp} /\ i{1} <= m{1} /\ + (forall (j : int), + i{1} <= j < m{1} => + dom HybridIROEager.mp{1} (xs{1}, j))). +wp; inline*. +rcondf{1} 3; first auto; smt(). +auto; smt(). +auto. +(* ! dom HybridIROEager.mp{2} (xs{2}, i{2}) *) +rcondf{2} 1; first auto. +sp; exists* i{1}; elim*=> i'. +conseq + (_ : + ={i, m, xs, bs, HybridIROEager.mp} /\ i{1} = i' /\ + i' + r = m{1} /\ size bs{1} = i' /\ cs{2} = [] /\ j{2} = 0 /\ + (forall (j : int), + i' <= j < i' + r => + ! dom HybridIROEager.mp{1} (xs{1}, j)) ==> + _). +progress; smt(gt0_r). +seq 1 2 : + (={m, xs} /\ i{2} = i' /\ i{1} = i' + r /\ bs{1} = bs{2} /\ + size bs{1} = i' + r /\ m{1} = i' + r /\ + (forall (k : int), + i' <= k < i' + r => + HybridIROEager.mp{1}.[(xs{1}, k)] = Some(nth true bs{1} k)) /\ + eager_eq_except xs{1} i' (i' + r) HybridIROEager.mp{1} HybridIROEager.mp{2}). +wp. +while + (={m, xs} /\ i{2} = i' /\ m{1} = i' + r /\ i' <= i{1} <= i' + r /\ + 0 <= j{2} <= r /\ i{1} - i' = j{2} /\ + size bs{1} = i{1} /\ bs{1} = bs{2} ++ cs{2} /\ + (forall (k : int), + i' <= k < i{1} => + HybridIROEager.mp{1}.[(xs{1}, k)] = Some(nth true bs{1} k)) /\ + (forall (k : int), + i{1} <= k < i' + r => + ! dom HybridIROEager.mp{1} (xs{1}, k)) /\ + eager_eq_except xs{1} i' (i' + r) HybridIROEager.mp{1} HybridIROEager.mp{2}). +inline*; rcondt{1} 3; first auto; smt(). +sp; wp; rnd; skip; progress. +smt(size_cat). smt(size_cat). smt(size_cat). +smt(size_rcons size_cat). smt(size_cat). +rewrite -cats1; smt(size_cat). +rewrite -2!cats1 catA; congr; congr. +by rewrite get_set_sameE oget_some. +rewrite nth_rcons /=. +case: (k = size (bs{2} ++ cs{2}))=> [-> /= | ne_k_size_bs_cat_cs]. +by rewrite get_set_sameE oget_some. +have -> /= : k < size(bs{2} ++ cs{2}) by smt(). +rewrite get_setE ne_k_size_bs_cat_cs /= /#. +rewrite -mem_fdom fdom_set in_fsetU1 mem_fdom negb_or. +have lt_sz_k : size (bs{2} ++ cs{2}) < k; smt(). +by apply eager_eq_except_upd1_eq_in. +smt(size_cat). smt(size_cat). +skip; progress; smt(gt0_r cats0 size_cat). +conseq + (_ : + ={xs, bs, m} /\ i{2} = i' /\ i{1} = i' + r /\ m{1} = i' + r /\ + size bs{1} = i' + r /\ + (forall (k : int), + i' <= k < i' + r => + HybridIROEager.mp{1}.[(xs{1}, k)] = Some (nth true bs{1} k)) /\ + eager_eq_except xs{1} i' (i' + r) + HybridIROEager.mp{1} HybridIROEager.mp{2} ==> + _)=> //. +while{2} + (={xs, bs, m} /\ i' <= i{2} <= i' + r /\ i{1} = i' + r /\ + m{1} = i' + r /\ size bs{1} = i' + r /\ + (forall (k : int), + i' <= k < i{2} => + HybridIROEager.mp{1}.[(xs{1}, k)] = HybridIROEager.mp{2}.[(xs{1}, k)]) /\ + (forall (k : int), + i{2} <= k < i' + r => + HybridIROEager.mp{1}.[(xs{1}, k)] = Some (nth true bs{1} k)) /\ + eager_eq_except xs{1} i' (i' + r) + HybridIROEager.mp{1} HybridIROEager.mp{2}) + (m{2} - i{2}). +progress; auto; progress; + [smt() | smt(gt0_r) | smt(get_setE) | smt() | + by apply eager_eq_except_upd2_eq_in | smt()]. +skip; progress; + [smt(gt0_r) | smt() | smt() | smt() | smt(eager_eq_except_maps_eq)]. +qed. + +(* module needed for applying transitivity tactic in connection + with BlockSponge.BIRO.IRO *) + +module BlockSpongeTrans = { + (* getting next block; assumes size bs = i *) + + proc next_block(x, i, bs) = { + var b; + + b <@ BlockSponge.BIRO.IRO.fill_in(x, i); + bs <- rcons bs b; + i <- i + 1; + return (bs, i); + } + + (* loop getting n blocks *) + + proc loop(n : int, xs : block list) : int * block list = { + var b : block; var i <- 0; var bs <- []; + while (i < n) { + b <@ BlockSponge.BIRO.IRO.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + return (i, bs); + } +}. + +(* use Program abstract theory of DList *) + +clone Program as Prog with + type t = bool, + op d = {0,1} +proof *. +(* nothing to be proved *) + +lemma PrLoopSnoc_sample &m (bs : bool list) : + Pr[Prog.LoopSnoc.sample(r) @ &m : res = bs] = + mu (dlist {0,1} r) (pred1 bs). +proof. +have -> : + Pr[Prog.LoopSnoc.sample(r) @ &m : res = bs] = + Pr[Prog.Sample.sample(r) @ &m : res = bs]. + byequiv=> //. + symmetry. + conseq (_ : ={n} ==> ={res})=> //. + apply Prog.Sample_LoopSnoc_eq. +apply (Prog.pr_Sample r &m bs). +qed. + +lemma iter_mul_one_half_pos (n : int) : + 0 < n => iter n (( * ) (1%r / 2%r)) 1%r = inv(2 ^ n)%r. +proof. +move=> gt0_n. +have -> /# // : + forall (n : int), + 0 <= n => 0 < n => iter n (( * ) (1%r / 2%r)) 1%r = inv (2 ^ n)%r. +elim=> [// | i ge0_i IH _]. +case: (i = 0)=> [-> /= | ne_i0]. +rewrite iter1 expr1 /#. +by rewrite iterS // IH 1:/# exprS // fromintM RField.invfM. +qed. + +(* module for adapting PrLoopSnoc_sample to block generation *) + +module BlockGen = { + proc loop() : block = { + var b : bool; var j : int; var cs : bool list; + j <- 0; cs <- []; + while (j < r) { + b <$ {0,1}; + cs <- rcons cs b; + j <- j + 1; + } + return mkblock cs; + } + + proc direct() : block = { + var w : block; + w <$ bdistr; + return w; + } +}. + +lemma BlockGen_loop_direct : + equiv[BlockGen.loop ~ BlockGen.direct : true ==> ={res}]. +proof. +bypr res{1} res{2}=> // &1 &2 w. +have -> : Pr[BlockGen.direct() @ &2 : res = w] = 1%r / (2 ^ r)%r. + byphoare=> //. + proc; rnd; skip; progress. + rewrite DBlock.dunifinE. + have -> : (transpose (=) w) = (pred1 w) by rewrite /pred1. + by rewrite DBlock.Support.enum_spec block_card. +have -> : + Pr[BlockGen.loop() @ &1 : res = w] = + Pr[Prog.LoopSnoc.sample(r) @ &1 : res = ofblock w]. + byequiv=> //; proc. + seq 2 2 : + (r = n{2} /\ j{1} = i{2} /\ j{1} = 0 /\ + cs{1} = l{2} /\ cs{1} = []); + first auto. + while + (r = n{2} /\ j{1} = i{2} /\ cs{1} = l{2} /\ j{1} <= r /\ + size cs{1} = j{1}). + wp; rnd; skip. + progress; smt(cats1 gt0_r size_rcons). + skip=> &m1 &m2 [# <- <- -> <- ->]. + split; first smt(gt0_r). + move=> cs j i ds not_lt_jr not_lt_ir [# _ eq_ji -> le_jr sz_cs_eq_j]. + have sz_ds_eq_r : size ds = r by smt(). + progress; [by rewrite ofblockK | by rewrite mkblockK]. +rewrite (PrLoopSnoc_sample &1 (ofblock w)). +rewrite dlist1E 1:ge0_r size_block /=. +have -> : + (fun (x : bool) => mu1 {0,1} x) = + (fun (x : bool) => 1%r / 2%r). +apply fun_ext=> x; by rewrite dbool1E. +by rewrite Bigreal.BRM.big_const count_predT size_block + iter_mul_one_half_pos 1:gt0_r. +qed. + +lemma HybridIROEagerTrans_BlockSpongeTrans_next_block (i2 : int) : + equiv + [HybridIROEagerTrans.next_block ~ BlockSpongeTrans.next_block : + i2 = i{2} /\ 0 <= i{2} /\ xs{1} = x{2} /\ i{1} = i{2} * r /\ + m{1} - i{1} = r /\ bs{1} = blocks2bits bs{2} /\ size bs{2} = i{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + res{1}.`1 = blocks2bits res{2}.`1 /\ + res{1}.`2 = res{2}.`2 * r /\ res{2}.`2 = i2 + 1 /\ + size res{2}.`1 = i2 + 1 /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}]. +proof. +transitivity + HybridIROEagerTrans.next_block_split + (={i, m, xs, bs, HybridIROEager.mp} /\ m{1} = i{1} + r /\ + size bs{1} = i{1} /\ + block_bits_dom_all_in_or_out xs{1} i{1} HybridIROEager.mp{1} ==> + ={res, HybridIROEager.mp}) + (i2 = i{2} /\ 0 <= i{2} /\ xs{1} = x{2} /\ i{1} = i{2} * r /\ + m{1} - i{1} = r /\ size bs{2} = i2 /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + res{1}.`1 = blocks2bits res{2}.`1 /\ + res{1}.`2 = res{2}.`2 * r /\ res{2}.`2 = i2 + 1 /\ + size res{2}.`1 = i2 + 1 /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +move=> |> &1 &2 ge0_i2 -> i1_eq_i2_tim_r m_min_i1_eq_r -> sz_bs_eq_i2 ei. +exists HybridIROEager.mp{1} (x{2}, i{1}, m{1}, blocks2bits bs{2})=> |>. +split; first smt(). +split; first smt(size_blocks2bits). +apply + (eager_inv_imp_block_bits_dom BlockSponge.BIRO.IRO.mp{2} + HybridIROEager.mp{1} x{2} i{1})=> //. +rewrite i1_eq_i2_tim_r mulr_ge0 // ge0_r. +rewrite i1_eq_i2_tim_r dvdz_mull dvdzz. +apply HybridIROEagerTrans_next_block_split. +proc=> /=; inline*; sp; wp. +case (dom BlockSponge.BIRO.IRO.mp{2} (x0{2}, n{2})). +(* dom BlockSponge.BIRO.IRO.mp{2} (x0{2}, n{2}) *) +rcondf{2} 1; first auto. +rcondt{1} 1; first auto; progress [-delta]. +have bb_all_in : + block_bits_all_in_dom x{m} (size bs{m} * r) HybridIROEager.mp{hr} + by apply (eager_inv_mem_dom1 BlockSponge.BIRO.IRO.mp{m}). +smt(gt0_r). +simplify; exists* i{1}; elim*=> i1; exists* bs{1}; elim*=> bs1. +conseq + (_ : + i1 = i{1} /\ 0 <= i2 /\ i1 = i2 * r /\ m{1} - i1 = r /\ + bs1 = bs{1} /\ size bs{2} = i2 /\ size bs1 = i1 /\ + bs1 = blocks2bits bs{2} /\ + dom BlockSponge.BIRO.IRO.mp{2} (xs{1}, i2) /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = + blocks2bits (rcons bs{2} (oget BlockSponge.BIRO.IRO.mp{2}.[(xs{1}, i2)])) /\ + i{1} = (i2 + 1) * r /\ size bs{2} = i2 /\ size bs{1} = (i2 + 1) * r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}); + [progress; smt(size_blocks2bits) | progress; by rewrite size_rcons | idtac]. +while{1} + (i1 <= i{1} <= m{1} /\ i1 = i2 * r /\ size bs{1} = i{1} /\ m{1} - i1 = r /\ + bs{1} = + bs1 ++ + take (i{1} - i1) + (ofblock (oget(BlockSponge.BIRO.IRO.mp{2}.[(xs{1}, i2)]))) /\ + dom BlockSponge.BIRO.IRO.mp{2} (xs{1}, i2) /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (m{1} - i{1}). +move=> &m z. +auto=> + |> &hr i2_tim_r_le_sz_bs sz_bs_le_m m_min_i2_tim_r_eq_r bs_eq + mem_blk_mp_xs_i2 ei sz_bs_lt_m. +split. split. split=> [| _]; smt(). split; first by rewrite -cats1 size_cat. +rewrite -cats1 {1}bs_eq -catA; congr. +have -> : size bs{hr} + 1 - i2 * r = size bs{hr} - i2 * r + 1 by algebra. +rewrite (take_nth false) 1:size_block; first smt(size_ge0). +rewrite -cats1; congr; congr. +have some_form_mp_hr_lookup_eq : + HybridIROEager.mp{hr}.[(xs{hr}, size bs{hr})] = + Some (nth false (ofblock (oget BlockSponge.BIRO.IRO.mp{m}.[(xs{hr}, i2)])) + (size bs{hr} - i2 * r)). + have [ei1 _] := ei. + have [_ ei1_xs_i2] := ei1 xs{hr} i2 mem_blk_mp_xs_i2. + by rewrite ei1_xs_i2 1:/#. +by rewrite some_form_mp_hr_lookup_eq oget_some. +smt(). +skip=> + &1 &2 + [# <- ge0_i2 i1_eq_i2_tim_r m_min_i1_eq_r <- sz_bs2_eq_i2 + sz_b2b_bs2_eq_i1 ->> mem_dom_mp2_xs_i2 ei]. +split. split. split=> [// | _]; rewrite i1_eq_i2_tim_r; smt(ge0_r). +split=> //. split; first smt(). split=> //. +split; first by rewrite /= take0 cats0. split=> //. +move=> bs1 i1'. +split=> [| not_i1'_lt_m]; first smt(). +move=> [# i1_le_i1' i1'_le_m _ sz_bs1_eq_i1' _ bs1_eq mem_mp2_xs_i2 _]. +split. +have i1'_eq_m : i1' = m{1} by smt(). +rewrite bs1_eq -cats1 blocks2bits_cat i1'_eq_m m_min_i1_eq_r blocks2bits_sing. +pose blk := (oget BlockSponge.BIRO.IRO.mp{2}.[(xs{1}, i2)]). +have -> : r = size (ofblock blk) by rewrite size_block. +by rewrite take_size. +split; smt(). +(* ! dom BlockSponge.BIRO.IRO.mp{2} (x0{2}, n{2}) *) +rcondt{2} 1; first auto. rcondf{1} 1; first auto; progress [-delta]. +have bb_all_not_in : + block_bits_all_out_dom x{m} (size bs{m} * r) HybridIROEager.mp{hr} + by apply (eager_inv_not_mem_dom1 BlockSponge.BIRO.IRO.mp{m}). +smt(gt0_r). +simplify. +conseq + (_ : + x0{2} = x{2} /\ n{2} = i{2} /\ i2 = i{2} /\ 0 <= i{2} /\ xs{1} = x{2} /\ + i{1} = i{2} * r /\ m{1} - i{1} = r /\ size bs{2} = i2 /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + _)=> //. +alias{2} 1 with w. +seq 3 1 : + (xs{1} = x0{2} /\ n{2} = i2 /\ i{2} = i2 /\ 0 <= i2 /\ + i{1} = i2 * r /\ m{1} - i{1} = r /\ size bs{2} = i2 /\ + size cs{1} = r /\ mkblock cs{1} = w{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +conseq (_ : true ==> cs{1} = ofblock w{2}); first + progress; [by rewrite size_block | by rewrite mkblockK]. +transitivity{2} + { w <@ BlockGen.loop(); } + (true ==> cs{1} = ofblock w{2}) + (true ==> ={w})=> //. +inline BlockGen.loop; sp; wp. +while (={j, cs} /\ 0 <= j{1} <= r /\ size cs{1} = j{1}). +wp; rnd; skip; progress; smt(size_ge0 size_rcons). +skip; progress. +smt(gt0_r). +have sz_cs_R_eq_r : size cs_R = r by smt(). +by rewrite ofblockK. +transitivity{2} + { w <@ BlockGen.direct(); } + (true ==> ={w}) + (true ==> ={w})=> //. +call BlockGen_loop_direct; auto. +inline BlockGen.direct; sim. +wp; simplify; sp; elim*=> bs_l. +exists* HybridIROEager.mp{1}; elim*=> mp1; exists* i{1}; elim*=> i1. +conseq + (_ : + xs{1} = x0{2} /\ 0 <= i2 /\ i{1} = i1 /\ i1 = i2 * r /\ + m{1} - i1 = r /\ + bs{1} = blocks2bits bs{2} ++ ofblock w{2} /\ size bs{2} = i2 /\ + size bs{1} = i1 + r /\ mp1 = HybridIROEager.mp{1} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = blocks2bits bs{2} ++ ofblock w{2} /\ + i{1} = (i2 + 1) * r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2}.[(x0{2}, i2) <- w{2}] + HybridIROEager.mp{1}). +progress; + [by rewrite ofblockK | rewrite size_cat size_blocks2bits /#]. +progress; + [by rewrite -cats1 blocks2bits_cat blocks2bits_sing get_set_sameE + oget_some ofblockK | + by rewrite size_rcons]. +while{1} + (0 <= i1 /\ m{1} - i1 = r /\ size bs{1} = i1 + r /\ + i1 <= i{1} <= m{1} /\ + eager_eq_except xs{1} i1 (i1 + r) mp1 HybridIROEager.mp{1} /\ + (forall (j : int), + i1 <= j < i{1} => + HybridIROEager.mp{1}.[(xs{1}, j)] = Some(nth false bs{1} j))) + (m{1} - i{1}). +progress; auto. +move=> + |> &hr ge0_i1 m_min_i1_eq_r sz_bs_eq_i1_plus_r il_le_i _ ee + mp_ran_eq lt_im. +split. split; first smt(). +split; first smt(eager_eq_except_upd2_eq_in). +move=> j i1_le_j j_lt_i_add1. +case: (i{hr} = j)=> [-> | ne_ij]. +rewrite get_setE /=; smt(nth_onth onth_nth). +rewrite get_setE. +have -> /= : (xs{hr}, j) <> (xs{hr}, i{hr}) by smt(). +rewrite mp_ran_eq /#. +smt(). +skip=> + &1 &2 [# -> ge0_i2 -> i1_eq_i2_tim_r m_min_i1_eq_r + bs1_eq sz_bs2_eq_i2 sz_bs1_eq_i1_add_r -> ei]. +have ge0_i1 : 0 <= i1 by rewrite i1_eq_i2_tim_r divr_ge0 // ge0_r. +split. split=> //. split; first smt(ge0_r). +split; first smt(). split; smt(ge0_r). +move=> mp_L i_L. +split; first smt(). +move=> not_i_L_lt_m [# _ _ _ i1_le_i_L i_L_le_m ee mp_L_ran_eq]. +split; first smt(). +split; first smt(). +apply (eager_invar_eq_except_upd1 BlockSponge.BIRO.IRO.mp{2} + HybridIROEager.mp{1} mp_L x0{2} i2 w{2})=> //. +by rewrite mulzDl /= -i1_eq_i2_tim_r. +move=> j j_ran. +rewrite mp_L_ran_eq 1:/#; congr; rewrite bs1_eq nth_cat. +have -> : size(blocks2bits bs{2}) = i2 * r + by rewrite size_blocks2bits /#. +have -> // : j < i2 * r = false by smt(). +qed. + +lemma HybridIROEagerTrans_BlockSpongeTrans_loop (n' : int) : + equiv + [HybridIROEagerTrans.loop ~ BlockSpongeTrans.loop : + ={xs, n} /\ n' = n{1} /\ 0 <= n' /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + res{1}.`1 = n' * r /\ res{2}.`1 = n' /\ + size res{2}.`2 = n' /\ res{1}.`2 = blocks2bits res{2}.`2 /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}]. +proof. +case (0 <= n'); last first=> [not_ge0_n' | ge0_n']. +proc=> /=; exfalso. +proc=> /=. +move: ge0_n'; elim n'=> [| n' ge0_n' IH]. +sp. rcondf{1} 1; auto. rcondf{2} 1; auto. +splitwhile{1} 3 : (i < (n - 1) * r); splitwhile{2} 3 : (i < n - 1). +seq 3 3 : + (={xs, n} /\ n{1} = n' + 1 /\ i{1} = n' * r /\ i{2} = n' /\ + size bs{2} = n' /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +conseq + (_ : + ={xs, n} /\ n' + 1 = n{1} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n' * r /\ i{2} = n' /\ size bs{2} = n' /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +transitivity{1} + { i <- 0; bs <- []; + while (i < n * r) { + b <@ HybridIROEager.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + } + (={xs, HybridIROEager.mp} /\ n{1} = n' + 1 /\ n{2} = n' ==> + ={bs, i, HybridIROEager.mp}) + (={xs} /\ n{1} = n' /\ n{2} = n' + 1 /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n' * r /\ i{2} = n' /\ size bs{2} = n' /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +progress; exists HybridIROEager.mp{1} n' xs{2}=> //. +while (={xs, i, bs, HybridIROEager.mp} /\ n{1} = n' + 1 /\ n{2} = n'). +wp. call (_ : ={HybridIROEager.mp}). if=> //; wp; rnd; auto. +skip; progress; smt(ge0_r). +auto; smt(). +transitivity{2} + { i <- 0; bs <- []; + while (i < n) { + b <@ BlockSponge.BIRO.IRO.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + } + (={xs, n} /\ n{1} = n' /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n' * r /\ i{2} = n' /\ size bs{2} = n' /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (={xs,BlockSponge.BIRO.IRO.mp} /\ n{1} = n' /\ n{2} = n' + 1 ==> + ={i, bs, BlockSponge.BIRO.IRO.mp})=> //. +progress; exists BlockSponge.BIRO.IRO.mp{2} n{1} xs{2}=> //. +conseq IH=> //. +while + (={xs, bs, i, BlockSponge.BIRO.IRO.mp} /\ n{1} = n' /\ n{2} = n' + 1). +wp. call (_ : ={BlockSponge.BIRO.IRO.mp}). if=> //; wp; rnd; auto. +auto; smt(). +auto; smt(). +unroll{2} 1. rcondt{2} 1; first auto; progress; smt(). +rcondf{2} 4. auto. call (_ : true). if=> //. auto. +transitivity{1} + { (bs, i) <@ HybridIROEagerTrans.next_block(xs, i, (n' + 1) * r, bs); } + (={xs, i, bs, HybridIROEager.mp} /\ n{1} = n' + 1 ==> + ={i, bs, HybridIROEager.mp}) + (={xs} /\ i{1} = n' * r /\ i{2} = n' /\ + size bs{2} = n' /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = (n' + 1) * r /\ i{2} = n' + 1 /\ size bs{2} = n' + 1 /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +progress; + exists HybridIROEager.mp{1} (blocks2bits bs{2}) (size bs{2} * r) xs{2}=> //. +inline HybridIROEagerTrans.next_block; sp; wp. +while + (xs{1} = xs0{2} /\ i{1} = i0{2} /\ n{1} = n' + 1 /\ + m{2} = (n' + 1) * r /\ bs{1} = bs0{2} /\ + ={HybridIROEager.mp}). +wp. call (_ : ={HybridIROEager.mp}). if=> //; wp; rnd; auto. +auto. auto. +transitivity{2} + { (bs, i) <@ BlockSpongeTrans.next_block(xs, i, bs); } + (={xs} /\ i{1} = n' * r /\ i{2} = n' /\ + size bs{2} = n' /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = (n' + 1) * r /\ i{2} = n' + 1 /\ size bs{2} = n' + 1 /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (={xs, bs, i, BlockSponge.BIRO.IRO.mp} ==> + ={xs, bs, i, BlockSponge.BIRO.IRO.mp})=> //. +progress; exists BlockSponge.BIRO.IRO.mp{2} bs{2} (size bs{2}) xs{2}=> //. +call (HybridIROEagerTrans_BlockSpongeTrans_next_block n'). +skip; progress; smt(). +inline BlockSpongeTrans.next_block. +wp; sp. call (_ : ={BlockSponge.BIRO.IRO.mp}). if=> //; wp; rnd; skip; smt(). +auto. +qed. + +lemma HybridIROEager_f_BlockIRO_f (n1 : int) (x2 : block list) : + equiv[HybridIROEager.f ~ BlockSponge.BIRO.IRO.f : + n1 = n{1} /\ x2 = x{2} /\ xs{1} = x{2} /\ + n{2} = (n{1} + r - 1) %/ r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} /\ + (valid_block x2 => + (n1 <= 0 => res{1} = [] /\ res{2} = []) /\ + (0 < n1 => + res{1} = take n1 (blocks2bits res{2}) /\ + size res{2} = (n1 + r - 1) %/ r)) /\ + (! valid_block x2 => res{1} = [] /\ res{2} = [])]. +proof. +proc=> /=. +seq 3 2 : + (n1 = n{1} /\ xs{1} = x{2} /\ x2 = x{2} /\ + n{2} = (n{1} + r - 1) %/ r /\ n{2} * r = m{1} /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +auto; progress. +if=> //. +case (n1 < 0). +rcondf{1} 1; first auto; progress; smt(). +rcondf{2} 1; first auto; progress; + by rewrite -lezNgt needed_blocks_non_pos ltzW. +rcondf{1} 1; first auto; progress; + by rewrite -lezNgt pmulr_lle0 1:gt0_r needed_blocks_non_pos ltzW. +by auto; progress; smt(needed_blocks0). +(* 0 <= n1 *) +conseq + (_ : + xs{1} = x{2} /\ n1 = n{1} /\ 0 <= n1 /\ n{2} = (n{1} + r - 1) %/ r /\ + n{2} * r = m{1} /\ n{1} <= m{1} /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = take n1 (blocks2bits bs{2}) /\ + size bs{2} = (n1 + r - 1) %/ r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +progress; [smt() | apply/needed_blocks_suff]. +move=> |> &1 &2 ? ? ? mp1 mp2 bs ? ? ?; + smt(size_eq0 needed_blocks0 take0). +splitwhile{1} 1 : i < (n1 %/ r) * r. splitwhile{2} 1 : i < n1 %/ r. +seq 1 1 : + (xs{1} = x{2} /\ n1 = n{1} /\ 0 <= n1 /\ n{2} = (n1 + r - 1) %/ r /\ + n{2} * r = m{1} /\ n{1} <= m{1} /\ i{1} = n1 %/ r * r /\ + i{2} = n1 %/ r /\ size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +(* we have zero or more blocks to add on the right, and + r times that number of bits to add on the left; + we will work up to applying HybridIROEagerTrans_BlockSpongeTrans_loop *) +conseq + (_ : + xs{1} = x{2} /\ n1 = n{1} /\ 0 <= n1 /\ n{2} = (n1 + r - 1) %/ r /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ + size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +transitivity{1} + { while (i < n1 %/ r * r) { + b <@ HybridIROEager.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + } + (={i, bs, xs, HybridIROEager.mp} /\ n1 = n{1} /\ 0 <= n1 ==> + ={i, bs, xs, HybridIROEager.mp}) + (xs{1} = x{2} /\ n1 = n{1} /\ 0 <= n1 /\ n{2} = (n1 + r - 1) %/ r /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ + size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +progress; exists HybridIROEager.mp{1} [] 0 n{1} x{2}=> //. +while (={i, bs, xs, HybridIROEager.mp} /\ n1 = n{1} /\ 0 <= n1). +wp. call (_ : ={HybridIROEager.mp}). if=> //; auto. +auto; progress; smt(leq_trunc_div ge0_r). +auto; progress; smt(leq_trunc_div ge0_r). +(transitivity{2} + { while (i < n1 %/ r) { + b <@ BlockSponge.BIRO.IRO.fill_in(x, i); + bs <- rcons bs b; + i <- i + 1; + } + } + (xs{1} = x{2} /\ n1 = n{1} /\ 0 <= n1 /\ n{2} = (n1 + r - 1) %/ r /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ size bs{2} = i{2} /\ + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (={i, x, bs, BlockSponge.BIRO.IRO.mp} /\ n{2} = (n1 + r - 1) %/ r ==> + ={i, x, bs, BlockSponge.BIRO.IRO.mp})=> //; + first progress; + exists BlockSponge.BIRO.IRO.mp{2} [] 0 ((n{1} + r - 1) %/ r) x{2}=> //); + first last. +while + (={i, x, bs, BlockSponge.BIRO.IRO.mp} /\ n{2} = (n1 + r - 1) %/ r). +wp. call (_ : ={BlockSponge.BIRO.IRO.mp}). if=> //; auto. +auto; progress; + have /# : n1 %/ r <= (n1 + r - 1) %/ r + by rewrite leq_div2r; smt(gt0_r). +auto; progress; + have /# : n1 %/ r <= (n1 + r - 1) %/ r + by rewrite leq_div2r; smt(gt0_r). +conseq + (_ : + xs{1} = x{2} /\ 0 <= n1 /\ + i{1} = 0 /\ i{2} = 0 /\ bs{1} = [] /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ + size bs{2} = n1 %/ r /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +transitivity{1} + { (i, bs) <@ HybridIROEagerTrans.loop(n1 %/ r, xs); } + (={xs, HybridIROEager.mp} /\ n{2} = n1 %/ r /\ i{1} = 0 /\ bs{1} = [] ==> + ={i, xs, bs, HybridIROEager.mp}) + (xs{1} = x{2} /\ 0 <= n1 /\ i{2} = 0 /\ bs{2} = [] /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ + size bs{2} = n1 %/ r /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +progress; exists HybridIROEager.mp{1} (n1 %/ r) x{2}=> //. +inline HybridIROEagerTrans.loop; sp; wp. +while + (={HybridIROEager.mp} /\ i{1} = i0{2} /\ bs{1} = bs0{2} /\ + xs{1} = xs0{2} /\ n0{2} = n1 %/ r). +wp. call (_ : ={HybridIROEager.mp}). if=> //; wp; rnd; auto. +auto. auto. +(transitivity{2} + { (i, bs) <@ BlockSpongeTrans.loop(n1 %/ r, x); } + (xs{1} = x{2} /\ 0 <= n1 /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + i{1} = n1 %/ r * r /\ i{2} = n1 %/ r /\ + size bs{2} = n1 %/ r /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (={x, BlockSponge.BIRO.IRO.mp} /\ i{2} = 0 /\ bs{2} = [] ==> + ={i, x, bs, BlockSponge.BIRO.IRO.mp})=> //; + first progress; exists BlockSponge.BIRO.IRO.mp{2} x{2}=> //); + last first. +inline BlockSpongeTrans.loop; sp; wp. +while + (={BlockSponge.BIRO.IRO.mp} /\ i0{1} = i{2} /\ n0{1} = n1 %/ r /\ + xs{1} = x{2} /\ bs0{1} = bs{2}). +wp. call (_ : ={BlockSponge.BIRO.IRO.mp}). if=> //; wp; rnd; auto. +auto. auto. +call (HybridIROEagerTrans_BlockSpongeTrans_loop (n1 %/ r)). +skip; progress; smt(divz_ge0 gt0_r). +(* either nothing more to do on either side, or a single block to add + on the right side, and less than r bits to add on the left side *) +conseq + (_ : + n1 = n{1} /\ 0 <= n1 /\ xs{1} = x{2} /\ n{2} = (n1 + r - 1) %/ r /\ + n{2} * r = m{1} /\ n{1} <= m{1} /\ i{1} = n1 %/ r * r /\ + i{2} = n1 %/ r /\ size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} /\ + (i{2} = n{2} \/ i{2} + 1 = n{2}) ==> + bs{1} = take n1 (blocks2bits bs{2}) /\ size bs{2} = n{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) => //. +progress; by apply/needed_blocks_rel_div_r. +case (i{2} = n{2}). (* so i{1} = n{1} and i{1} = m{1} *) +rcondf{2} 1; first auto; progress; smt(). +rcondf{1} 1; first auto; progress; smt(). +rcondf{1} 1; first auto; progress; smt(). +auto=> |> &1 &2 ? ? sz_eq ? ? need_blks_eq. +split. +have -> : n1 = size (blocks2bits bs{2}) + by rewrite size_blocks2bits sz_eq -mulzC divzK 1:needed_blocks_eq_div_r. +by rewrite take_size. +by rewrite sz_eq need_blks_eq. +(* i{2} <> n{2}, so i{2} + 1 = n{2}, m{1} - i{1} = r and i{1} <= m{1} *) +rcondt{2} 1; first auto; progress; smt(). +rcondf{2} 4. +auto; call (_ : true); [if=> //; auto; progress; smt() | auto; smt()]. +conseq + (_ : + n1 = n{1} /\ 0 <= n1 /\ xs{1} = x{2} /\ 0 <= i{2} /\ i{1} = i{2} * r /\ + n{1} <= m{1} /\ m{1} - i{1} = r /\ i{1} <= n{1} /\ + bs{1} = blocks2bits bs{2} /\ size bs{1} = i{1} /\ size bs{2} = i{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = take n1 (blocks2bits bs{2}) /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + _ + (_ : size bs = n - 1 ==> size bs = n)=> //. +progress; smt(divz_ge0 gt0_r lez_floor size_blocks2bits). +wp. call (_ : true). auto. skip; smt(size_rcons). +transitivity{1} + { while (i < m) { + b <@ HybridIROEager.fill_in(xs, i); + bs <- rcons bs b; + i <- i + 1; + } + } + (={bs, m, i, xs, HybridIROEager.mp} /\ n1 = n{1} /\ i{1} <= n1 /\ + n1 <= m{1} /\ size bs{1} = i{1} ==> + ={HybridIROEager.mp} /\ bs{1} = take n1 bs{2}) + (xs{1} = x{2} /\ 0 <= i{2} /\ i{1} = i{2} * r /\ m{1} - i{1} = r /\ + size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +progress; + exists HybridIROEager.mp{1} (blocks2bits bs{2}) + (size bs{2} * r) m{1} x{2}=> //. +progress; smt(take_cat). +splitwhile{2} 1 : i < n1. +seq 1 1 : + (={HybridIROEager.mp, xs, bs, i, m} /\ i{1} = n1 /\ n1 <= m{1} /\ + size bs{1} = n1). +while + (={HybridIROEager.mp, xs, bs, i, m} /\ n{1} = n1 /\ n1 <= m{1} /\ + i{1} <= n1 /\ size bs{1} = i{1}). +wp; call (_ : ={HybridIROEager.mp}); first if => //; wp; rnd; auto. +skip; smt(size_rcons). +skip; smt(). +while + (={HybridIROEager.mp, xs, i, m} /\ n1 <= m{1} /\ + n1 <= i{1} <= m{1} /\ n1 <= size bs{2} /\ + bs{1} = take n1 bs{2}). +wp; call (_ : ={HybridIROEager.mp}); first if => //; wp; rnd; auto. +skip; progress; + [smt() | smt() | smt(size_rcons) | + rewrite -cats1 take_cat; + smt(size_rcons take_oversize cats1 cats0)]. +skip; smt(take_size). +(* now we can use HybridIROEagerTrans_BlockSpongeTrans_next_block *) +transitivity{1} + { (bs, i) <@ HybridIROEagerTrans.next_block(xs, i, m, bs); + } + (={i, m, xs, bs, HybridIROEager.mp} ==> + ={i, m, xs, bs, HybridIROEager.mp}) + (xs{1} = x{2} /\ 0 <= i{2} /\ i{1} = i{2} * r /\ m{1} - i{1} = r /\ + size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1})=> //. +progress [-delta]; + exists HybridIROEager.mp{1} (blocks2bits bs{2}) (size bs{2} * r) + m{1} x{2}=> //. +inline HybridIROEagerTrans.next_block; sim. +(transitivity{2} + { (bs, i) <@ BlockSpongeTrans.next_block(x, i, bs); + } + (xs{1} = x{2} /\ 0 <= i{2} /\ i{1} = i{2} * r /\ m{1} - i{1} = r /\ + size bs{2} = i{2} /\ bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + bs{1} = blocks2bits bs{2} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}) + (={bs, i, x, BlockSponge.BIRO.IRO.mp} ==> + ={bs, i, x, BlockSponge.BIRO.IRO.mp})=> //; + first progress [-delta]; + exists BlockSponge.BIRO.IRO.mp{2} bs{2} (size bs{2}) x{2}=> //); + last first. +inline BlockSpongeTrans.next_block; sim. +exists* i{2}; elim*=> i2. +call (HybridIROEagerTrans_BlockSpongeTrans_next_block i2). +auto. +qed. + +lemma HybridIROEager_BlockIRO_f : + equiv[LowerHybridIRO(HybridIROEager).f ~ BlockSponge.BIRO.IRO.f : + xs{1} = x{2} /\ ={n} /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + ={res} /\ eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}]. +proof. +transitivity + HybridIROEager.f + (={xs, HybridIROEager.mp} /\ n{2} = n{1} * r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + res{1} = bits2blocks res{2} /\ ={HybridIROEager.mp}) + (xs{1} = x{2} /\ n{1} = n{2} * r /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1} ==> + res{1} = (blocks2bits res{2}) /\ + eager_invar BlockSponge.BIRO.IRO.mp{2} HybridIROEager.mp{1}). +move=> |> &1 &2 ? n_eq inv. +exists HybridIROEager.mp{1} BlockSponge.BIRO.IRO.mp{2} (xs{1}, n{1} * r). +move=> |>; by rewrite n_eq. +progress; apply blocks2bitsK. +by conseq Lower_HybridIROEager_f=> |> &1 &2 ? -> ?. +exists* n{1}; elim*=> n1; exists* xs{1}; elim*=> xs'. +conseq (HybridIROEager_f_BlockIRO_f n1 xs')=> //. +move=> |> &1 &2 ? -> inv; by rewrite needed_blocks_prod_r. +move=> |> &1 &2 ? n1_eq ? res1 res2 ? ? ? vb_imp not_vb_imp. +case: (valid_block xs{1})=> [vb_xs1 | not_vb_xs1]. +have [le0_n1_imp gt0_n1_imp] := vb_imp vb_xs1. +case: (n{1} <= 0)=> [le0_n1 /# | not_le0_n1]. +have gt0_n1 : 0 < n{1} by smt(). +have [-> sz_res2] := gt0_n1_imp gt0_n1. +have -> : n{1} = size(blocks2bits res2) + by rewrite size_blocks2bits sz_res2 n1_eq + needed_blocks_prod_r mulzC. +by rewrite take_size. +by have [-> ->] := not_vb_imp not_vb_xs1. +qed. + +end HybridIRO. + +(* now we use HybridIRO to prove the main result *) + +section. + +declare module BlockSim <: BlockSponge.SIMULATOR {-IRO, -BlockSponge.BIRO.IRO}. +declare module Dist <: DISTINGUISHER {-Perm, -BlockSim, -IRO, -BlockSponge.BIRO.IRO}. + +local clone HybridIRO as HIRO. + +(* working toward the Real side of the main result *) + +local lemma Sponge_Raise_BlockSponge_f : + equiv[Sponge(Perm).f ~ RaiseFun(BlockSponge.Sponge(Perm)).f : + ={bs, n, glob Perm} ==> ={res, glob Perm}]. +proof. +proc; inline BlockSponge.Sponge(Perm).f. +conseq (_ : ={bs, n, glob Perm} ==> _)=> //. +swap{2} [3..5] -2. +seq 4 4 : + (={n, glob Perm, sa, sc, i} /\ xs{1} = xs0{2} /\ z{1} = [] /\ z{2} = [] /\ + valid_block xs0{2}). +auto; progress; apply valid_pad2blocks. +rcondt{2} 2; auto. swap{2} 1 1. +seq 1 1 : + (={n, glob Perm, sa, sc, i} /\ xs{1} = xs0{2} /\ z{1} = [] /\ z{2} = []). +while (={glob Perm, sa, sc, i} /\ xs{1} = xs0{2} /\ z{1} = [] /\ z{2} = []). +wp. call (_ : ={glob Perm}). sim. auto. auto. +seq 0 1 : + (={n, glob Perm, sa, sc, i} /\ blocks2bits z{2} = z{1} /\ + n0{2} = (n{1} + r - 1) %/ r); first auto. +while (={n, glob Perm, i, sa, sc} /\ blocks2bits z{2} = z{1} /\ + n0{2} = (n{1} + r - 1) %/ r). +case (i{1} + 1 < (n{1} + r - 1) %/ r). +rcondt{1} 3; first auto. rcondt{2} 3; first auto. +call (_ : ={glob Perm}); first sim. +auto; progress; by rewrite -cats1 blocks2bits_cat blocks2bits_sing. +rcondf{1} 3; first auto. rcondf{2} 3; first auto. +auto; progress; by rewrite -cats1 blocks2bits_cat blocks2bits_sing. +auto. +qed. + +(* the Real side of main result *) +local lemma RealIndif_Sponge_BlockSponge &m : + Pr[RealIndif(Sponge, Perm, Dist).main() @ &m : res] = + Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, LowerDist(Dist)).main() @ &m : res]. +proof. +byequiv=> //; proc. +seq 2 2 : (={glob Dist, glob Perm}); first sim. +call (_ : ={glob Perm}); first 2 sim. +conseq Sponge_Raise_BlockSponge_f=> //. +auto. +qed. + +(* working toward the Ideal side of the main result *) + +(* first step of Ideal side: express in terms of Experiment and + HIRO.HybridIROLazy *) + +local lemma Ideal_IRO_Experiment_HybridLazy &m : + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res] = + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROLazy), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROLazy)), + Dist).main() @ &m : res]. +proof. +byequiv=> //; proc. +seq 2 2 : + (={glob Dist, glob BlockSim} /\ IRO.mp{1} = empty /\ + HIRO.HybridIROLazy.mp{2} = empty). +inline*; wp; call (_ : true); auto. +call + (_ : + ={glob Dist, glob BlockSim} /\ + IRO.mp{1} = empty /\ HIRO.HybridIROLazy.mp{2} = empty ==> + ={res}). +proc + (={glob BlockSim} /\ + HIRO.lazy_invar IRO.mp{1} HIRO.HybridIROLazy.mp{2})=> //. +progress [-delta]; apply HIRO.lazy_invar0. +proc (HIRO.lazy_invar IRO.mp{1} HIRO.HybridIROLazy.mp{2})=> //; + apply HIRO.LowerFun_IRO_HybridIROLazy_f. +proc (HIRO.lazy_invar IRO.mp{1} HIRO.HybridIROLazy.mp{2})=> //; + apply HIRO.LowerFun_IRO_HybridIROLazy_f. +by conseq HIRO.IRO_RaiseHybridIRO_HybridIROLazy_f. +auto. +qed. + +(* working toward middle step of Ideal side: using Experiment, and + taking HIRO.HybridIROLazy to HIRO.HybridIROEager + + we will employ HIRO.HybridIROExper_Lazy_Eager *) + +(* make a Hybrid IRO distinguisher from BlockSim and Dist (HI.f is + used by BlockSim, and HI.g is used by HIRO.RaiseHybridIRO; + HI.init is unused -- see the SIMULATOR module type) *) + +local module (HybridIRODist : HIRO.HYBRID_IRO_DIST) (HI : HIRO.HYBRID_IRO) = { + proc distinguish() : bool = { + var b : bool; + BlockSim(HIRO.LowerHybridIRO(HI)).init(); + b <@ + Dist(HIRO.RaiseHybridIRO(HI), + BlockSim(HIRO.LowerHybridIRO(HI))).distinguish(); + return b; + } +}. + +(* initial bridging step *) + +local lemma Experiment_HybridIROExper_Lazy &m : + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROLazy), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROLazy)), + Dist).main() @ &m : res] = + Pr[HIRO.HybridIROExper(HIRO.HybridIROLazy, HybridIRODist).main() @ &m : res]. +proof. +byequiv=> //; proc; inline*. +seq 2 2 : (={glob Dist, glob BlockSim, HIRO.HybridIROLazy.mp}). +swap{2} 1 1; wp; call (_ : true); auto. +sim. +qed. + +(* final bridging step *) + +local lemma HybridIROExper_Experiment_Eager &m : + Pr[HIRO.HybridIROExper(HIRO.HybridIROEager, HybridIRODist).main() @ + &m : res] = + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROEager), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROEager)), + Dist).main() @ &m : res]. +proof. +byequiv=> //; proc; inline*. +seq 2 2 : (={glob Dist, glob BlockSim, HIRO.HybridIROEager.mp}). +swap{2} 1 1; wp; call (_ : true); auto. +sim. +qed. + +(* middle step of Ideal side: using Experiment, and taking HIRO.HybridIROLazy + to HIRO.HybridIROEager *) + +local lemma Experiment_Hybrid_Lazy_Eager &m : + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROLazy), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROLazy)), + Dist).main() @ &m : res] = + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROEager), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROEager)), + Dist).main() @ &m : res]. +proof. +by rewrite (Experiment_HybridIROExper_Lazy &m) + (HIRO.HybridIROExper_Lazy_Eager HybridIRODist &m) + (HybridIROExper_Experiment_Eager &m). +qed. + +(* working toward last step of Ideal side *) + +local lemma RaiseHybridIRO_HybridIROEager_RaiseFun_BlockIRO_f : + equiv[HIRO.RaiseHybridIRO(HIRO.HybridIROEager).f ~ + RaiseFun(BlockSponge.BIRO.IRO).f : + ={bs, n} /\ ={glob BlockSim} /\ + HIRO.eager_invar BlockSponge.BIRO.IRO.mp{2} HIRO.HybridIROEager.mp{1} ==> + ={res} /\ ={glob BlockSim} /\ + HIRO.eager_invar BlockSponge.BIRO.IRO.mp{2} HIRO.HybridIROEager.mp{1}]. +proof. +proc=> /=. +exists* n{1}; elim*=> n'. +exists* (pad2blocks bs{2}); elim*=> xs2. +call (HIRO.HybridIROEager_f_BlockIRO_f n' xs2). +skip=> |> &1 &2 ? res1 res2 mp1 mp2 ? vb_imp not_vb_imp. +case: (valid_block (pad2blocks bs{2}))=> [vb | not_vb]. +have [le0_n2_imp gt0_n2_imp] := vb_imp vb. +case: (n' <= 0)=> [le0_n' /# | not_le0_n']. +have gt0_n' : 0 < n' by smt(). +by have [-> _] := gt0_n2_imp gt0_n'. +have [-> ->] := not_vb_imp not_vb; by rewrite blocks2bits_nil. +qed. + +(* last step of Ideal side: express in terms of Experiment and + HIRO.HybridIROEager *) + +local lemma Experiment_HybridEager_Ideal_BlockIRO &m : + Pr[Experiment + (HIRO.RaiseHybridIRO(HIRO.HybridIROEager), + BlockSim(HIRO.LowerHybridIRO(HIRO.HybridIROEager)), + Dist).main() @ &m : res] = + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main () @ &m : res]. +proof. +byequiv=> //; proc. +seq 2 2 : + (={glob Dist, glob BlockSim} /\ HIRO.HybridIROEager.mp{1} = empty /\ + BlockSponge.BIRO.IRO.mp{2} = empty). +inline*; wp; call (_ : true); auto. +call + (_ : + ={glob Dist, glob BlockSim} /\ + HIRO.HybridIROEager.mp{1} = empty /\ BlockSponge.BIRO.IRO.mp{2} = empty ==> + ={res}). +proc + (={glob BlockSim} /\ + HIRO.eager_invar BlockSponge.BIRO.IRO.mp{2} HIRO.HybridIROEager.mp{1})=> //. +progress [-delta]; apply HIRO.eager_invar0. +proc (HIRO.eager_invar BlockSponge.BIRO.IRO.mp{2} + HIRO.HybridIROEager.mp{1})=> //; + conseq HIRO.HybridIROEager_BlockIRO_f=> //. +proc (HIRO.eager_invar BlockSponge.BIRO.IRO.mp{2} + HIRO.HybridIROEager.mp{1})=> //; + conseq HIRO.HybridIROEager_BlockIRO_f=> //. +conseq RaiseHybridIRO_HybridIROEager_RaiseFun_BlockIRO_f=> //. +auto. +qed. + +(* the Ideal side of main result *) + +local lemma IdealIndif_IRO_BlockIRO &m : + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res] = + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main () @ &m : res]. +proof. +by rewrite (Ideal_IRO_Experiment_HybridLazy &m) + (Experiment_Hybrid_Lazy_Eager &m) + (Experiment_HybridEager_Ideal_BlockIRO &m). +qed. + +lemma conclu &m : + `|Pr[RealIndif(Sponge, Perm, Dist).main() @ &m : res] - + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res]| = + `|Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, LowerDist(Dist)).main() @ &m : res] - + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main() @ &m : res]|. +proof. +by rewrite (RealIndif_Sponge_BlockSponge &m) (IdealIndif_IRO_BlockIRO &m). +qed. + +end section. + +(*----------------------------- Conclusion -----------------------------*) + +lemma conclusion + (BlockSim <: BlockSponge.SIMULATOR{-IRO, -BlockSponge.BIRO.IRO}) + (Dist <: DISTINGUISHER{-Perm, -BlockSim, -IRO, -BlockSponge.BIRO.IRO}) + &m : + `|Pr[RealIndif(Sponge, Perm, Dist).main() @ &m : res] - + Pr[IdealIndif(IRO, RaiseSim(BlockSim), Dist).main() @ &m : res]| = + `|Pr[BlockSponge.RealIndif + (BlockSponge.Sponge, Perm, LowerDist(Dist)).main() @ &m : res] - + Pr[BlockSponge.IdealIndif + (BlockSponge.BIRO.IRO, BlockSim, LowerDist(Dist)).main() @ &m : res]|. +proof. by apply (conclu BlockSim Dist &m). qed. diff --git a/sha3/proof/smart_counter/ConcreteF.eca b/sha3/proof/smart_counter/ConcreteF.eca new file mode 100644 index 0000000..5c510f2 --- /dev/null +++ b/sha3/proof/smart_counter/ConcreteF.eca @@ -0,0 +1,442 @@ +require import Core Int Real StdOrder Ring Distr. +require import List FSet SmtMap Common SLCommon DProd Dexcepted. + +(*...*) import Capacity IntID IntOrder RealOrder. + +require (*..*) Strong_RP_RF. + +module PF = { + var m, mi: (state,state) fmap + + proc init(): unit = { + m <- empty; + mi <- empty; + } + + proc f(x : state): state = { + var y1, y2; + + if (x \notin m) { + y1 <$ bdistr; + y2 <$ cdistr; + m.[x] <- (y1,y2); + mi.[(y1,y2)] <- x; + } + return oget m.[x]; + } + + proc fi(x : state): state = { + var y1, y2; + + if (x \notin mi) { + y1 <$ bdistr; + y2 <$ cdistr; + mi.[x] <- (y1,y2); + m.[(y1,y2)] <- x; + } + return oget mi.[x]; + } + +}. + +module CF(D:DISTINGUISHER) = Indif(SqueezelessSponge(PF), PF, D). + +section. + declare module D <: DISTINGUISHER {-Perm, -C, -PF, -Redo}. + + declare axiom D_ll (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}): + islossless P.f => islossless P.fi => islossless F.f => + islossless D(F, P).distinguish. + + local module GReal' = Indif(FC(SqueezelessSponge(Perm)), PC(Perm), D). + + local clone import Strong_RP_RF as Switching with + type D <- state, + op uD <- dstate, + type K <- unit, + op dK <- (MUnit.dunit<:unit> tt), + op q <- max_size + proof *. + realize ge0_q by smt w=max_ge0. + realize uD_uf_fu. + split. + case=> [x y]; rewrite supp_dprod /=. + rewrite Block.DBlock.supp_dunifin Capacity.DCapacity.supp_dunifin/=. + smt(dprod1E Block.DBlock.dunifin_funi Capacity.DCapacity.dunifin_funi). + split. + smt(dprod_ll Block.DBlock.dunifin_ll Capacity.DCapacity.dunifin_ll). + apply/dprod_fu. + rewrite Block.DBlock.dunifin_fu. + by rewrite Capacity.DCapacity.dunifin_fu. + qed. + realize dK_ll. + by rewrite /is_lossless MUnit.dunit_ll. + qed. + + (* TODO move this *) + lemma size_behead (l : 'a list) : l <> [] => size (behead l) = size l - 1. + proof. by case l=> // ?? /=; ring. qed. + + local module (D': PRPSec.Distinguisher) (P' : PRPSec.SPRP_Oracles) = { + proc distinguish () : bool = { + var b : bool; + Redo.init(); + b <@ DRestr(D, SqueezelessSponge(P'), P').distinguish(); + return b;} + }. + + local lemma DoubleBounding (P <: PRPSec.PRP {-D, -C, -DBounder, -Redo}) &m: + Pr[PRPSec.IND(P,D').main() @ &m: res] + = Pr[PRPSec.IND(P,DBounder(D')).main() @ &m: res]. + proof. + byequiv=> //=; proc; inline *. + wp. + call (_: ={glob C, glob P, glob Redo} + /\ all_prefixes Redo.prefixes{2} + /\ Redo.prefixes{2}.[[]] = Some (b0,c0) + /\ (forall x, x \in C.queries{2} => x \in Redo.prefixes{2}) + /\ prefix_inv C.queries{2} Redo.prefixes{2} + /\ DBounder.FBounder.c{2} = C.c{2}). + + proc; sp; if=> //=; inline *. + rcondt{2} 3; 1: by auto=> /#. + by wp; call (_: true); auto. + + proc; sp; if=> //=; inline *. + rcondt{2} 3; 1: by auto=> /#. + by wp; call (_: true); auto. + + proc; sp; if=> //=; inline *;1:if;auto. + + splitwhile{1}5:take (i+1) p \in Redo.prefixes. + splitwhile{2}5:take (i+1) p \in Redo.prefixes. + + alias{1}1 pref = Redo.prefixes;alias{2}1 pref = Redo.prefixes;sp 1 1=>/=. + alias{1}1 query = C.queries;alias{2}1 query = C.queries;sp 1 1=>/=. + conseq(:_==> ={sa, Redo.prefixes, glob P, i, C.c} + /\ all_prefixes Redo.prefixes{2} + /\ (forall x, x \in query{2} => x \in Redo.prefixes{2}) + /\ i{1} = size bs{1} + /\ Redo.prefixes{1}.[take i{1} bs{1}] = Some (sa{1},sc{1}) + /\ (forall y, y \in pref{1} => pref{1}.[y] = Redo.prefixes{1}.[y]) + /\ (forall y, y \in Redo.prefixes{1} <=> (y \in pref{1} \/ + (exists j, 0 <= j <= i{1} /\ y = take j bs{1}))) + /\ DBounder.FBounder.c{2} = C.c{2} - size bs{1} + i{1}). + + auto=> />; progress. + + smt(domE mem_set get_setE take_size cat_take_drop). + + smt(domE mem_set get_setE take_size cat_take_drop). + + smt(domE mem_set get_setE take_size cat_take_drop). + + smt(domE mem_set get_setE take_size cat_take_drop). + + by move: H15=> /H11 []; smt(domE mem_set get_setE take_size cat_take_drop). + smt(domE mem_set get_setE take_size cat_take_drop). + while( ={sa, Redo.prefixes, glob P, i, C.c, p, sc} /\ p{1} = bs{1} + /\ all_prefixes Redo.prefixes{2} + /\ Redo.prefixes{2}.[[]] = Some (b0, c0) + /\ (forall x, x \in query{2} => x \in Redo.prefixes{2}) + /\ (i{1} < size p{1} => ! take (i{1} + 1) p{1} \in Redo.prefixes{1}) + /\ 0 <= prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) <= i{1} <= size bs{1} + /\ C.c{1} <= max_size + /\ Redo.prefixes{1}.[take i{1} bs{1}] = Some (sa{1},sc{1}) + /\ (forall y, y \in pref{1} => pref{1}.[y] = Redo.prefixes{1}.[y]) + /\ (forall y, y \in Redo.prefixes{1} <=> (y \in pref{1} \/ + (exists j, 0 <= j <= i{1} /\ y = take j bs{1}))) + /\ DBounder.FBounder.c{2} = C.c{2} - size bs{1} + i{1}). + + if; auto; 1:smt(domE). + sp;rcondt{2}1;1:auto=>/#;auto;1:call(:true);auto;progress. + * move=>x;rewrite mem_set=>[][|-> j]; 1:smt(mem_set). + case(0 <= j)=>hj0;last first. + + by rewrite (@take_le0 j)1:/# domE get_setE H0 /#. + by rewrite take_take /min; case: (j < i{2} + 1); rewrite mem_set //= /#. + * smt(mem_set take_take domE get_setE oget_some). + * smt(mem_set take_take domE get_setE oget_some). + * rewrite mem_set negb_or H9 negb_or/=negb_exists/=. + have htake:take (i{2} + 1) bs{1} = take (i{2} + 1) (take (i{2} + 1 + 1) bs{1}); + smt(take_take size_take). + * rewrite/#. + * rewrite/#. + * smt(mem_set take_take domE get_setE oget_some). + * smt(mem_set take_take domE get_setE oget_some). + * smt(mem_set take_take domE get_setE oget_some). + * smt(mem_set take_take domE get_setE oget_some). + * smt(mem_set take_take domE get_setE oget_some). + sp; + conseq(:_==> ={sa, Redo.prefixes, glob P, i, C.c, p, sc, bs} /\ p{1} = bs{1} + /\ Redo.prefixes{2} = pref{2} + /\ (forall x, x \in query{2} => x \in Redo.prefixes{2}) + /\ C.c{1} <= max_size + /\ i{1} = prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) + /\ Redo.prefixes{1}.[take i{1} bs{1}] = Some (sa{1},sc{1}) + /\ DBounder.FBounder.c{2} = C.c{2} - size bs{1} + + prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))))=> />. + progress. + + rewrite -negP. + move: H9; rewrite (prefix_exchange _ Redo.prefixes{2} _)=> //= H9. + by rewrite -mem_fdom memE; apply/prefix_lt_size=> /#. + + exact/prefix_ge0. + + exact/prefix_sizel. + + case: H9=> //= - [j] [#] H42 H72. + have ->: j = min j (prefix bs{2} (get_max_prefix bs{2} (elems (fdom C.queries{2})))) by smt(). + rewrite -(take_take bs{2} j (prefix bs{2} (get_max_prefix bs{2} (elems (fdom C.queries{2}))))). + by move=> ->; rewrite H domE //= H8. + + smt(). + alias{2} 1 k = DBounder.FBounder.c;sp; + conseq(:_==> ={sa, Redo.prefixes, glob P, i, C.c, p, sc, bs} /\ p{1} = bs{1} + /\ Redo.prefixes{2} = pref{2} + /\ (forall x, x \in query{2} => x \in Redo.prefixes{2}) + /\ C.c{2} <= max_size + /\ i{1} = prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) + /\ Redo.prefixes{1}.[take i{1} bs{1}] = Some (sa{1},sc{1}) + /\ DBounder.FBounder.c{2} = k{2});1:progress=>/#. + while( ={sa, Redo.prefixes, glob P, i, C.c, p, sc, bs, C.queries} /\ p{1} = bs{1} + /\ Redo.prefixes{2} = pref{2} + /\ (forall x, x \in query{2} => x \in Redo.prefixes{2}) + /\ prefix_inv C.queries{2} Redo.prefixes{2} + /\ all_prefixes Redo.prefixes{2} + /\ C.c{2} <= max_size + /\ 0 <= i{1} <= prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) + /\ (forall j, 0 <= j <= prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) + => take j bs{2} \in Redo.prefixes{1}) + /\ Redo.prefixes{1}.[take i{1} bs{1}] = Some (sa{1},sc{1}) + /\ DBounder.FBounder.c{2} = k{2}). + + rcondt{1}1;2:rcondt{2}1;auto;progress. + * by smt(). + * by rewrite(@prefix_exchange _ _ bs{2} H0 H1)all_take_in//=/#. + * smt(domE). + auto;progress. smt(prefix_ge0). + + apply/take_get_max_prefix2=> //=. + + by exists []; rewrite domE H0. + by rewrite-(@prefix_exchange _ _ _ H2 H). + * smt(domE take0). + * smt(@Prefix). + auto; call(: true); auto=> />. + (** TODO: send to smt with cast into infinite maps **) + do!split. + + rewrite/all_prefixes; smt(mem_set mem_empty). + + exact/get_set_sameE. + + smt(mem_empty mem_set). + + smt(mem_empty mem_set get_setE). + move=> ?; split=> [|_]. + + smt(mem_set mem_empty). + smt(mem_empty mem_set). + qed. + + local clone import ProdSampling with + type t1 <- block, + type t2 <- capacity. + + lemma Real_Concrete &m : + Pr[GReal(D).main()@ &m: res /\ C.c <= max_size] <= + Pr[CF(DRestr(D)).main()@ &m: res] + (max_size ^ 2 - max_size)%r / 2%r * mu dstate (pred1 witness). + proof. + have ->: + Pr[RealIndif(SqueezelessSponge,PC(Perm),D).main()@ &m: + res /\ C.c <= max_size] = Pr[GReal'.main()@ &m: res/\ C.c <= max_size]. + + byequiv=>//;proc;inline *; + call (_: ={C.c, glob Perm, Redo.prefixes} + /\ prefix_inv C.queries{2} Redo.prefixes{1} + /\ Redo.prefixes{1}.[[]] = Some (b0, c0) + /\ all_prefixes Redo.prefixes{1}); + last first. + + auto;smt(mem_empty mem_set get_setE oget_some). + + by proc; inline*; sp; if; auto. + + by proc; inline*; sp; if; auto. + proc; inline *; wp; sp. + if{2};sp;wp;last first. + + conseq(:_==> sa{1} = (oget Redo.prefixes{1}.[take i{1} p{1}]).`1 + /\ i{1} = size p{1} + /\ Redo.prefixes{1}.[[]] = Some (b0, c0) + /\ ={Perm.m, Perm.mi, Redo.prefixes, C.c});1:smt(take_size). + + while{1}( ={Perm.m, Perm.mi, Redo.prefixes, C.c} + /\ p{1} \in C.queries{2} + /\ prefix_inv C.queries{2} Redo.prefixes{1} + /\ 0 <= i{1} <= size p{1} + /\ Redo.prefixes{1}.[[]] = Some (b0, c0) + /\ (sa{1},sc{1}) = oget Redo.prefixes{1}.[take i{1} p{1}] + /\ all_prefixes Redo.prefixes{1})(size p{1} - i{1}). + + auto;sp;rcondt 1;auto;smt(excepted_lossless). + by auto;smt(size_ge0 take0 take_size). + + splitwhile{1} 1 : take (i+1) p \in Redo.prefixes; + splitwhile{2} 1 : take (i+1) p \in Redo.prefixes. + + alias{1}1 pref = Redo.prefixes;alias{2}1 pref = Redo.prefixes;sp 1 1=>/=. + alias{2}1 query = C.queries;sp 0 1=>/=. + + conseq(:_==> ={sa,Perm.m,Perm.mi,Redo.prefixes,i,p} + /\ C.c{1} = C.c{2} - size p{2} + i{2} + /\ i{2} = size p{2} + /\ Redo.prefixes{2}.[take i{2} p{2}] = Some (sa{2}, sc{2}) + /\ (forall l, l \in pref{2} => pref{2}.[l] = Redo.prefixes{2}.[l]) + /\ (forall j, 0 <= j <= i{2} => take j p{2} \in Redo.prefixes{2}) + /\ (forall l, l \in Redo.prefixes{2} => + l \in pref{2} \/ (exists j, 0 <= j <= i{2} /\ l = take j p{2}))); + progress. + * by rewrite/#. + * move:H3 H7;rewrite take_size mem_set get_setE;case(bs0 = bs{2})=>//=[->|]h. + * by rewrite h oget_some/=. + * move:H=>[H []];progress. + by rewrite -H4; move: (H3 _ H9 (size bs0)); rewrite take_size //= H. + * smt(mem_set take_size oget_some get_setE domE take_oversize take_le0). + * elim: (H6 _ H10). + + elim: H=> /> _ _ /(_ bs0 i0 H9) h /h [] l2 hl2. + by exists l2; rewrite mem_set hl2. + by move=> [j] [] hj ->; exists (drop j bs{2}); rewrite cat_take_drop mem_set. + * smt(mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop). + * smt(mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop). + while(={sa,sc,Perm.m,Perm.mi,Redo.prefixes,i,p} + /\ C.c{1} = C.c{2} - size p{2} + i{2} + /\ all_prefixes Redo.prefixes{2} + /\ all_prefixes pref{2} + /\ prefix_inv C.queries{2} pref{2} + /\ prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) <= i{2} <= size p{2} + /\ Redo.prefixes{2}.[take i{2} p{2}] = Some (sa{2}, sc{2}) + /\ (forall l, l \in pref{2} => pref{2}.[l] = Redo.prefixes{2}.[l]) + /\ (forall j, 0 <= j <= i{2} => take j p{2} \in Redo.prefixes{2}) + /\ (forall l, l \in Redo.prefixes{2} => + l \in pref{2} \/ (exists j, 0 <= j <= i{2} /\ l = take j p{2}))). + + rcondf{1}1;2:rcondf{2}1;..2:auto;progress. + * have:=H7 (take (i{m0}+1) p{m0}). + case((take (i{m0} + 1) p{m0} \in Redo.prefixes{m0}))=>//=_. + rewrite negb_or negb_exists/=;progress. + + by rewrite -mem_fdom memE prefix_lt_size//=-(@prefix_exchange _ _ p{m0} H1 H0)//=/#. + case(0<=a<=i{m0})=>//=ha;smt(size_take). + * have:=H7 (take (i{hr}+1) p{hr}). + case((take (i{hr} + 1) p{hr} \in Redo.prefixes{hr}))=>//=_. + rewrite negb_or negb_exists/=;progress. + + by rewrite -mem_fdom memE prefix_lt_size//=-(@prefix_exchange _ _ p{hr} H1 H0)//=/#. + case(0<=a<=i{hr})=>//=ha;smt(size_take). + + sp;auto;if;auto;progress. + * rewrite/#. + * move=>x;rewrite mem_set=>[][|h];1: + smt(mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop). + rewrite h=>j;rewrite take_take /min. + case(j//=hij. + case: (0 <= j)=> hj //=. + + by rewrite mem_set; left; apply/H6=> /#. + rewrite mem_set (take_le0 j) 1:/#; left. + by rewrite -(take0 (take i{2} p{2})); apply/H/domE; rewrite H4. + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE). + * rewrite!get_setE/=. + have/#: !take (i{2} + 1) p{2} \in pref{2}. + by rewrite -mem_fdom memE prefix_lt_size//=-(@prefix_exchange _ _ _ H1 H0)//=/#. + * rewrite get_set_sameE !oget_some. + have: take (i{2} + 1) p{2} \notin Redo.prefixes{2}. + + move: (H7 (take (i{2} + 1) p{2})); case: (take (i{2} + 1) p{2} \in Redo.prefixes{2})=> //= _. + rewrite negb_or negb_exists //=; split. + + rewrite -mem_fdom memE; apply/prefix_lt_size. + + by rewrite -(prefix_exchange C.queries{2}) // /#. + by rewrite -(prefix_exchange C.queries{2}) // /#. + smt(size_take). + rewrite domE=> /= H728; rewrite get_set_neqE 2:H5 //. + have /H5:= H13. + by apply/contraLR=> /= ->>; move: H13; rewrite domE H728=> ->. + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE). + * smt(). + * move=>x;rewrite mem_set =>[][|h];1: + smt(mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop). + rewrite h=>j;rewrite take_take /min. + case(j//=hij. + case(0 <= j)=> //= hj. + + by rewrite mem_set H6 /#. + rewrite (take_le0 j) 1:/# mem_set. + have:= (H (take i{2} p{2}) _ 0). + + by rewrite domE H4. + by rewrite take0=> ->. + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + * by rewrite!get_setE. + * rewrite !get_setE//=. + have /#: !take (i{2} + 1) p{2} \in pref{2}. + by rewrite -mem_fdom memE prefix_lt_size//=-(@prefix_exchange _ _ _ H1 H0)//=/#. + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + * smt(prefix_lt_size mem_set take_size oget_some get_setE domE take_oversize take_le0 take_take cat_take_drop memE mem_fdom). + conseq(:_==> ={sa,sc,Perm.m,Perm.mi,Redo.prefixes,i,p} + /\ C.c{1} = C.c{2} - size p{2} + i{2} + /\ pref{2} = Redo.prefixes{2} + /\ all_prefixes pref{2} + /\ prefix_inv C.queries{2} pref{2} + /\ prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) = i{2} + /\ Redo.prefixes{2}.[take i{2} p{2}] = Some (sa{2}, sc{2}));1: + smt(prefix_sizel take_get_max_prefix2 domE prefix_exchange). + + while( ={sa,sc,Perm.m,Perm.mi,Redo.prefixes,i,p} + /\ C.c{1} = C.c{2} - size p{2} + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) + /\ pref{2} = Redo.prefixes{2} + /\ all_prefixes pref{2} + /\ prefix_inv C.queries{2} pref{2} + /\ 0 <= i{2} <= prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) + /\ Redo.prefixes{2}.[take i{2} p{2}] = Some (sa{2}, sc{2})). + + rcondt{1}1;2:rcondt{2}1;auto;progress. + * rewrite/#. search get_max_prefix (<=) take mem. + * rewrite(@prefix_inv_leq _ _ _ _ _ _ H H7 H0)//= 1:/#. + have :=H0=>[][h1 [h2 h3]]. + have :=h3 _ _ _ H7;last smt(memE mem_fdom). + smt(size_eq0 size_take). + * smt(domE). + auto;progress. + * rewrite/#. + * smt(prefix_ge0). + * smt(take0). + * smt(prefix_sizel @Prefix memE). + * smt(prefix_sizel @Prefix memE). + + have p_ll := f_ll _. + + by move=> [b c]; rewrite supp_dprod /= Block.DBlock.dunifin_fu Capacity.DCapacity.dunifin_fu. + have pi_ll := fi_ll _. + + by move=> [b c]; rewrite supp_dprod /= Block.DBlock.dunifin_fu Capacity.DCapacity.dunifin_fu. + have f_ll : islossless SqueezelessSponge(Perm).f. + + proc; while true (size p - i)=> //=. + * by move=> z; wp;if;auto; 2:call (p_ll); auto=>/#. + by auto; smt w=size_ge0. + apply (@ler_trans _ _ _ + (Pr_restr Perm SqueezelessSponge D p_ll pi_ll f_ll D_ll &m)). + have ->: Pr[Indif(SqueezelessSponge(Perm), Perm, DRestr(D)).main() @ &m: res] + = Pr[PRPSec.IND(PRPi.PRPi,DBounder(D')).main() @ &m: res]. + + rewrite -(DoubleBounding PRPi.PRPi &m). + byequiv=> //=; proc; inline *; sim (_: ={m,mi}(Perm,PRPi.PRPi) /\ ={glob C}). + (** * by proc; if=> //=; auto. **) + * proc. if. + + move=> &1 &2 [#] <<- _ _ -> _. (** FIXME: the two instances of PRPi.PRPi.mi{2} appear to not be the same value; one of them in an ill-formed term **) smt(). + + auto=> /#. + + auto=> /#. + by proc; if=> //=; auto=> /#. + have ->: Pr[CF(DRestr(D)).main() @ &m: res] + = Pr[PRPSec.IND(ARP,DBounder(D')).main() @ &m: res]. + + rewrite -(DoubleBounding ARP &m). + byequiv=> //=; proc; inline *; sim (_: ={m,mi}(PF,ARP)). + * proc; if=> //=; auto; conseq (_: true ==> (y1,y2){1} = x{2})=> //=. + transitivity{1} { (y1,y2) <@ S.sample2(bdistr,cdistr); } + (true ==> ={y1,y2}) + (true ==> (y1,y2){1} = x{2})=> //=. + - by inline *; auto. + transitivity{2} { x <@ S.sample(bdistr,cdistr); } + (true ==> (y1,y2){1} = x{2}) + (true ==> ={x})=> //=. + - by symmetry; call sample_sample2; skip=> /> []. + by inline *; auto. + proc; if=> //=; auto; conseq (_: true ==> (y1,y2){1} = y{2})=> //=. + transitivity{1} { (y1,y2) <@ S.sample2(bdistr,cdistr); } + (true ==> ={y1,y2}) + (true ==> (y1,y2){1} = y{2})=> //=. + - by inline *; auto. + transitivity{2} { y <@ S.sample(bdistr,cdistr); } + (true ==> (y1,y2){1} = y{2}) + (true ==> ={y})=> //=. + - by symmetry; call sample_sample2; skip=> /> []. + by inline *; auto. + have:= Conclusion D' &m _. + + move=> O O_f_ll O_fi_ll. + proc;inline*;sp;wp; call (_: true)=> //=. + + apply D_ll. + + by proc; inline*; sp; if=> //=; auto; call O_f_ll; auto. + + by proc; inline*; sp; if=> //=; auto; call O_fi_ll; auto. + + proc; inline *; sp; if=> //=; auto; if; auto. + while true (size p - i);auto. + * sp; if; auto; 2:call O_f_ll; auto=> /#. + by auto; smt w=size_ge0. + smt(). + qed. + +end section. diff --git a/sha3/proof/smart_counter/Gcol.eca b/sha3/proof/smart_counter/Gcol.eca new file mode 100644 index 0000000..e5b2fe0 --- /dev/null +++ b/sha3/proof/smart_counter/Gcol.eca @@ -0,0 +1,358 @@ +pragma -oldip. +require import Core Int Real StdOrder Ring StdBigop. +require import List FSet SmtMap Common SLCommon FelTactic Mu_mem. +require import DProd Dexcepted PROM. +(*...*) import Capacity IntID IntOrder Bigreal RealOrder BRA. + +require (*..*) Handle. + +clone export Handle as Handle0. +import ROhandle. +(* -------------------------------------------------------------------------- *) + + (* TODO: move this *) + lemma c_gt0r : 0%r < (2^c)%r. + proof. by rewrite lt_fromint; apply/IntOrder.expr_gt0. qed. + + lemma c_ge0r : 0%r <= (2^c)%r. + proof. by apply /ltrW/c_gt0r. qed. + + lemma eps_ge0 : 0%r <= (2 * max_size)%r / (2 ^ c)%r. + proof. + apply divr_ge0;1:by rewrite le_fromint;smt ml=0 w=max_ge0. + by apply c_ge0r. + qed. + +section PROOF. + declare module D <: DISTINGUISHER{-C, -PF, -G1}. + + declare axiom D_ll (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}): + islossless P.f => islossless P.fi => + islossless F.f => islossless D(F, P).distinguish. + + local module Gcol = { + + var count : int + + proc sample_c () = { + var c <- c0; + if (card (image fst (frng FRO.m)) <= 2*max_size /\ + count < max_size) { + c <$ cdistr; + G1.bcol <- G1.bcol \/ mem (image fst (frng FRO.m)) c; + count <- count + 1; + } + + return c; + } + + module M = { + + proc f(p : block list): block = { + var sa, sa', sc; + var h, i, counter <- 0; + sa <- b0; + while (i < size p ) { + if ((sa +^ nth witness p i, h) \in G1.mh) { + (sa, h) <- oget G1.mh.[(sa +^ nth witness p i, h)]; + } else { + if (counter < size p - prefix p (get_max_prefix p (elems (fdom C.queries)))) { + sc <@ sample_c(); + sa' <@ F.RO.get(take (i+1) p); + sa <- sa +^ nth witness p i; + G1.mh.[(sa,h)] <- (sa', G1.chandle); + G1.mhi.[(sa',G1.chandle)] <- (sa, h); + (sa,h) <- (sa',G1.chandle); + FRO.m.[G1.chandle] <- (sc,Unknown); + G1.chandle <- G1.chandle + 1; + counter <- counter + 1; + } + } + i <- i + 1; + } + sa <@ F.RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2, hy2, hx2; + + if (x \notin G1.m) { + y <- (b0,c0); + if (!(rng FRO.m (x.`2, Known))) { + FRO.m.[G1.chandle] <- (x.`2, Known); + G1.chandle <- G1.chandle + 1; + } + hx2 <- oget (hinvK FRO.m x.`2); + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + y1 <@ F.RO.get (rcons p (v +^ x.`1)); + y2 <@ sample_c(); + } else { + y1 <$ bdistr; + y2 <@ sample_c(); + } + y <- (y1,y2); + if ((x.`1, hx2) \in G1.mh /\ + in_dom_with FRO.m (oget G1.mh.[(x.`1,hx2)]).`2 Unknown) { + hy2 <- (oget G1.mh.[(x.`1, hx2)]).`2; + y <- (y.`1, (oget FRO.m.[hy2]).`1); + FRO.m.[hy2] <- (y.`2, Known); + G1.m.[x] <- y; + G1.mi.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + FRO.m.[hy2] <- (y.`2, Known); + G1.m.[x] <- y; + G1.mh.[(x.`1, hx2)] <- (y.`1, hy2); + G1.mi.[y] <- x; + G1.mhi.[(y.`1, hy2)] <- (x.`1, hx2); + } + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + G1.paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget G1.m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2, hx2, hy2; + + if (x \notin G1.mi) { + y <- (b0,c0); + if (!(rng FRO.m (x.`2, Known))) { + FRO.m.[G1.chandle] <- (x.`2, Known); + G1.chandle <- G1.chandle + 1; + } + hx2 <- oget (hinvK FRO.m x.`2); + y1 <$ bdistr; + y2 <@ sample_c(); + y <- (y1,y2); + if ((x.`1, hx2) \in G1.mhi /\ + in_dom_with FRO.m (oget G1.mhi.[(x.`1,hx2)]).`2 Unknown) { + (y1,hy2) <- oget G1.mhi.[(x.`1, hx2)]; + y <- (y.`1, (oget FRO.m.[hy2]).`1); + FRO.m.[hy2] <- (y.`2, Known); + G1.mi.[x] <- y; + G1.m.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + FRO.m.[hy2] <- (y.`2, Known); + G1.mi.[x] <- y; + G1.mhi.[(x.`1, hx2)] <- (y.`1, hy2); + G1.m.[y] <- x; + G1.mh.[(y.`1, hy2)] <- (x.`1, hx2); + } + } else { + y <- oget G1.mi.[x]; + } + return y; + } + + } + + proc main(): bool = { + var b; + + F.RO.m <- empty; + G1.m <- empty; + G1.mi <- empty; + G1.mh <- empty; + G1.mhi <- empty; + G1.bcol <- false; + + FRO.m <- empty.[0 <- (c0, Known)]; + G1.paths <- empty.[c0 <- ([<:block>],b0)]; + G1.chandle <- 1; + count <- 0; + b <@ DRestr(D,M,S).distinguish(); + return b; + } + }. + + lemma card_rng_set (m:('a,'b)fmap) x y: card(frng m.[x<-y]) <= card(frng m) + 1. + proof. + have: frng m.[x <- y] \subset frng m `|` fset1 y. + + move=> b; rewrite in_fsetU1 2!mem_frng 2!rngE /= => [] [] a. + rewrite get_setE; case: (a = x) =>[->>|hax] //= hmab; left. + by exists a. + move=> /subset_leq_fcard; rewrite fcardU fcard1; smt(fcard_ge0). + qed. + + lemma hinv_image handles c: + hinv handles c <> None => + mem (image fst (frng handles)) c. + proof. + case: (hinv handles c) (hinvP handles c)=>//= h[f] Heq. + rewrite imageP;exists (c,f)=>@/fst/=. + by rewrite mem_frng rngE /=; exists (oget (Some h)). + qed. + + local lemma Pr_col &m : + Pr[Gcol.main()@&m : G1.bcol /\ Gcol.count <= max_size] <= + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + fel 10 Gcol.count (fun x=> (2*max_size)%r / (2^c)%r) + max_size G1.bcol + [Gcol.sample_c : (card (image fst (frng FRO.m)) <= 2*max_size /\ Gcol.count < max_size)]=>//;2:by auto. + + rewrite /felsum Bigreal.sumr_const count_predT size_range. + apply ler_wpmul2r;1:by apply eps_ge0. + by rewrite le_fromint;smt ml=0 w=max_ge0. + + proc;sp;if;2:by hoare=>//??;apply eps_ge0. + wp. + rnd (mem (image fst (frng FRO.m)));skip;progress;2:smt ml=0. + have->:=(Mu_mem.mu_mem (image fst (frng FRO.m{hr})) cdistr (1%r/(2^c)%r) _). + + move=>x _; rewrite DCapacity.dunifin1E;do !congr;smt(@Capacity). + apply ler_wpmul2r;2:by rewrite le_fromint. + by apply divr_ge0=>//;apply /c_ge0r. + + move=>ci;proc;rcondt 2;auto=>/#. + move=> b c;proc;sp;if;auto;smt ml=0. + qed. + + local equiv G1_col : G1(DRestr(D)).main ~ Gcol.main : + ={glob D} ==> (G1.bcol{1} => G1.bcol{2}) /\ Gcol.count{2} <= max_size. + proof. + proc;inline*;wp. + call (_: ={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m,C.c,C.queries}/\ + (G1.bcol{1} => G1.bcol{2}) /\ + (card (frng FRO.m) <= 2*C.c + 1 /\ + Gcol.count <= C.c <= max_size){2}). + + proc;sp 1 1;if=>//;inline G1(DRestr(D)).S.f Gcol.S.f;swap -3. + sp;if;1,3:auto=>/#;swap{1}[1..2]3;sp 1 1. + seq 5 5 : (={x0, y0, F.RO.m, G1.mi, G1.paths, G1.m, G1.mhi, G1.chandle, + G1.mh, FRO.m, C.c, C.queries} + /\ (G1.bcol{1} => G1.bcol{2}) + /\ card (frng FRO.m{2}) <= 2 * C.c{2} + 1 + /\ Gcol.count{2} <= C.c{2} <= max_size );last by if;auto. + seq 2 2 : (={x0, hx2, F.RO.m, G1.mi, G1.paths, G1.m, G1.mhi, G1.chandle, + G1.mh, FRO.m, C.c, C.queries} + /\ (G1.bcol{1} => G1.bcol{2}) + /\ card (frng FRO.m{2}) <= 2 * C.c{2} + /\ Gcol.count{2} + 1 <= C.c{2} <= max_size). + + by if; auto=> /> &1 &2; smt(card_rng_set). + if;1:auto. + - inline Gcol.sample_c;rcondt{2}4. + * auto;inline*;auto;progress. + + by have/#:=fcard_image_leq (fun (p : capacity * flag) => p.`1) (frng FRO.m{hr}). + rewrite/#. + seq 3 4 : (={x0, p, v, y1, hx2, F.RO.m, G1.mi, G1.paths, G1.m, G1.mhi, + G1.chandle, G1.mh, FRO.m, C.c, C.queries} + /\ (G1.bcol{1} => G1.bcol{2}) + /\ card (frng FRO.m{2}) <= 2 * C.c{2} + /\ Gcol.count{2} + 1 <= C.c{2} <= max_size + /\ (x0{1}.`2 \in G1.paths{1}) + /\ y2{1} = c{2});1: by inline*;auto. + sp 1 4;if;auto;progress. + + by have->:=(H H6). + + smt(card_rng_set). + + case:H5=>/=[h|H_hinv];1: by have->:=H h. + have:=hinvP FRO.m{2} c{2};rewrite H_hinv/=imageP/==>[][]f H_f. + by right; exists (c{2}, f)=> //=; rewrite mem_frng rngE/= /#. + smt(card_rng_set). + inline Gcol.sample_c;rcondt{2}3. + * auto;progress. + + by have/#:=fcard_image_leq (fun (p : capacity * flag) => p.`1) (frng FRO.m{hr}). + rewrite/#. + seq 2 3 : (={x0, y1, hx2, F.RO.m, G1.mi, G1.paths, G1.m, G1.mhi, + G1.chandle, G1.mh, FRO.m, C.c, C.queries} + /\ (G1.bcol{1} => G1.bcol{2}) + /\ card (frng FRO.m{2}) <= 2 * C.c{2} + /\ Gcol.count{2} + 1 <= C.c{2} <= max_size + /\ ! (x0{1}.`2 \in G1.paths{1}) + /\ y2{1} = c{2});1: by auto. + sp 1 4;if;auto;progress. + + by have->:=(H H6). + + smt(card_rng_set). + + case:H5=>/=[h|H_hinv];1: by have->:=H h. + have:= hinvP FRO.m{2} c{2}. + rewrite H_hinv /= imageP /= => [] [] f H_f. + by right; exists (c{2},f); rewrite mem_frng rngE /=; exists (oget (hinv FRO.m{2} c{2})). + smt(card_rng_set). + + + proc;sp 1 1;if=>//. + inline G1(DRestr(D)).S.fi Gcol.S.fi;swap-3. + seq 2 2 : (={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m, + C.c,C.queries,x0} /\ + (G1.bcol{1} => G1.bcol{2}) /\ + (card (frng FRO.m) + 2 <= 2*C.c + 1 /\ + Gcol.count + 1 <= C.c <= max_size){2});1:by auto=>/#. + if=>//;last by auto=>/#. + seq 3 3:(={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m, + C.c,C.queries,x0,hx2} /\ + (G1.bcol{1} => G1.bcol{2}) /\ + (card (frng FRO.m) + 1 <= 2 * C.c + 1 /\ + Gcol.count + 1 <= C.c <= max_size){2}). + + sp 1 1;if;auto;smt ml=0 w=card_rng_set. + seq 3 3: + (={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m, + C.c,C.queries,x0,hx2,y0,y1,y2} /\ y0{1} = (y1,y2){1} /\ + ((G1.bcol\/hinv FRO.m y0.`2 <> None){1} => G1.bcol{2}) /\ + (card (frng FRO.m) + 1 <= 2 * C.c + 1 /\ + Gcol.count <= C.c <= max_size){2});2:by auto;smt w=card_rng_set. + inline Gcol.sample_c. + rcondt{2}3. + + by auto;progress;have /#:= fcard_image_leq fst (frng FRO.m{hr}). +(* BUG: auto=> /> ?? Himp _ _ _ ?_?_ [/Himp->// | H]. marche pas ???? *) + auto=> /> ?? Himp _ _ _ ?_?_ [/Himp->// | X];right;apply hinv_image=> //. + + + proc;sp 1 1;if=>//;2:auto;sp;if=>//;swap 1;wp. + inline G1(DRestr(D)).M.f Gcol.M.f;sp;wp. + seq 1 1: + (={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m,C.c, + C.queries,b,p,h,i,sa,bs,counter} /\ i{1}=size p{2} /\ p{2} = bs{2} /\ + (G1.bcol{1} => G1.bcol{2}) /\ + (0 <= counter{2} <= size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1})))) /\ + card (frng FRO.m{2}) <= 2 * (C.c{2} + counter{2}) + 1 /\ + Gcol.count{2} + size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1}))) + - counter{2} <= C.c{2} + size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1}))) + <= max_size); + last by inline*;auto;smt(size_ge0 prefix_sizel). + while + (={F.RO.m,G1.mi,G1.paths,G1.m,G1.mhi,G1.chandle,G1.mh,FRO.m,C.c,b, + p,h,i,sa,counter,C.queries} /\ (0 <= i <= size p){1} /\ + (G1.bcol{1} => G1.bcol{2}) /\ + (0 <= counter{2} <= size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1})))) /\ + card (frng FRO.m{2}) <= 2 * (C.c{2} + counter{2}) + 1 /\ + Gcol.count{2} + size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1}))) + - counter{2} <= C.c{2} + size p{2} - + prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{1}))) + <= max_size);last first. + + by auto=> />; smt(size_ge0 prefix_sizel prefix_ge0). + if=>//;auto;1:smt ml=0 w=size_ge0. + if=>//;2:auto;2:smt(size_ge0 prefix_sizel). + auto;call (_: ={F.RO.m})=>/=;1:by sim. + inline *;rcondt{2} 2. + + auto;progress. + - apply(StdOrder.IntOrder.ler_trans _ _ _ (fcard_image_leq fst (frng FRO.m{hr})))=>/#. + smt(size_ge0 prefix_sizel). + auto;smt ml=0 w=(hinv_image card_rng_set). + auto;progress;3:by smt ml=0. + + rewrite -(add0z 1) -{2}fcards0<:capacity*flag> -(frng0<:int,_>). + exact/card_rng_set/max_ge0. + by apply max_ge0. + qed. + + lemma Pr_G1col &m: + Pr[G1(DRestr(D)).main() @ &m : G1.bcol] <= max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + apply (ler_trans Pr[Gcol.main()@&m : G1.bcol /\ Gcol.count <= max_size]). + + byequiv G1_col=> //#. + apply (Pr_col &m). + qed. + +end section PROOF. + + diff --git a/sha3/proof/smart_counter/Gconcl.ec b/sha3/proof/smart_counter/Gconcl.ec new file mode 100644 index 0000000..f114324 --- /dev/null +++ b/sha3/proof/smart_counter/Gconcl.ec @@ -0,0 +1,389 @@ +pragma -oldip. +require import Core Int Real StdOrder Ring StdBigop. +require import List FSet SmtMap Common SLCommon FelTactic Mu_mem. +require import DProd Dexcepted PROM. +(*...*) import Capacity IntID IntOrder Bigreal RealOrder BRA. + +require (*..*) Gext. + +module IF = { + proc init = F.RO.init + proc f = F.RO.get +}. + +module S(F : DFUNCTIONALITY) = { + var m, mi : smap + var paths : (capacity, block list * block) fmap + + proc init() = { + m <- empty; + mi <- empty; + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + paths <- empty.[c0 <- ([<:block>],b0)]; + } + + proc f(x : state): state = { + var p, v, y, y1, y2; + + if (x \notin m) { + if (x.`2 \in paths) { + (p, v) <- oget paths.[x.`2]; + y1 <@ F.f (rcons p (v +^ x.`1));} else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + m.[x] <- y; + mi.[y] <- x; + if (x.`2 \in paths) { + (p,v) <- oget paths.[x.`2]; + paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2; + + if (x \notin mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + mi.[x] <- y; + m.[y] <- x; + } else { + y <- oget mi.[x]; + } + return y; + } + +}. + +section. + +declare module D <: DISTINGUISHER{-C, -Perm, -F.RO, -F.FRO, -S, -Redo}. +local clone import Gext as Gext0. + + +local module G3(RO:F.RO) = { + + module M = { + + proc f(p : block list): block = { + var sa, sa'; + var h, i, counter <- 0; + sa <- b0; + while (i < size p ) { + if ((sa +^ nth witness p i, h) \in G1.mh) { + RO.sample(take (i+1) p); + (sa, h) <- oget G1.mh.[(sa +^ nth witness p i, h)]; + } else { + if (counter < size p - prefix p (get_max_prefix p (elems (fdom C.queries)))) { + RRO.sample(G1.chandle); + sa' <@ RO.get(take (i+1) p); + sa <- sa +^ nth witness p i; + G1.mh.[(sa,h)] <- (sa', G1.chandle); + G1.mhi.[(sa',G1.chandle)] <- (sa, h); + (sa,h) <- (sa',G1.chandle); + G1.chandle <- G1.chandle + 1; + counter <- counter + 1; + } else { + RO.sample(take (i+1) p); + } + } + i <- i + 1; + } + sa <@ RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2, hy2, hx2, handles_,t; + + if (x \notin G1.m) { + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + y1 <@ RO.get (rcons p (v +^ x.`1)); + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1, y2); + handles_ <@ RRO.allKnown(); + if (!rng handles_ x.`2) { + RRO.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ RRO.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + t <@ RRO.queried((oget G1.mh.[(x.`1,hx2)]).`2, Unknown); + if ((x.`1, hx2) \in G1.mh /\ t) { + hy2 <- (oget G1.mh.[(x.`1, hx2)]).`2; + FRO.m.[hy2] <- (y2,Known); + G1.m.[x] <- y; + G1.mi.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + RRO.set(hy2, y.`2); + G1.m.[x] <- y; + G1.mh.[(x.`1, hx2)] <- (y.`1, hy2); + G1.mi.[y] <- x; + G1.mhi.[(y.`1, hy2)] <- (x.`1, hx2); + } + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + G1.paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget G1.m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2, hx2, hy2, handles_, t; + + if (x \notin G1.mi) { + handles_ <@ RRO.allKnown(); + if (!rng handles_ x.`2) { + RRO.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ RRO.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + t <@ RRO.queried((oget G1.mhi.[(x.`1,hx2)]).`2, Unknown); + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + if ((x.`1, hx2) \in G1.mhi /\ t) { + (y1,hy2) <- oget G1.mhi.[(x.`1, hx2)]; + FRO.m.[hy2] <- (y2,Known); + G1.mi.[x] <- y; + G1.m.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + RRO.set(hy2, y.`2); + G1.mi.[x] <- y; + G1.mhi.[(x.`1, hx2)] <- (y.`1, hy2); + G1.m.[y] <- x; + G1.mh.[(y.`1, hy2)] <- (x.`1, hx2); + } + } else { + y <- oget G1.mi.[x]; + } + return y; + } + + } + + proc distinguish(): bool = { + var b; + + RO.init(); + G1.m <- empty; + G1.mi <- empty; + G1.mh <- empty; + G1.mhi <- empty; + + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + RRO.init(); + RRO.set(0,c0); + G1.paths <- empty.[c0 <- ([<:block>],b0)]; + G1.chandle <- 1; + b <@ DRestr(D,M,S).distinguish(); + return b; + } +}. + +local equiv G2_G3: Eager(G2(DRestr(D))).main2 ~ G3(F.LRO).distinguish : ={glob D} ==> ={res}. +proof. + proc;wp;call{1} RRO_resample_ll;inline *;wp. + call (_: ={FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries}); last by auto. + + + proc;sp;if=> //;sim. + call (_: ={FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries});2:by auto. + if=> //;2:by sim. + swap{1} [3..7] -2;swap{2} [4..8] -3. + seq 5 5:(={hx2,t,x,FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries} /\ + (t = in_dom_with FRO.m (oget G1.mh.[(x.`1, hx2)]).`2 Unknown){1}); + 1:by inline *;auto. + seq 3 4:(={y,x,FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries}); + 2:by sim. + if=>//. + + seq 2 2:(={y1,hx2,t,x,FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries} + /\ (t = in_dom_with FRO.m (oget G1.mh.[(x.`1, hx2)]).`2 Unknown){1}). + + by inline *;auto=> /> ? _;rewrite Block.DWord.bdistr_ll. + case (((x.`1, hx2) \in G1.mh /\ t){1}); + [rcondt{1} 3;2:rcondt{2} 3| rcondf{1} 3;2:rcondf{2} 3]; + 1,2,4,5:(by move=>?;conseq (_:true);auto);2:by sim. + inline *;rcondt{1} 6;1:by auto=>/>. + by auto => /> *; rewrite !get_setE. + case (((x.`1, hx2) \in G1.mh /\ t){1}); + [rcondt{1} 4;2:rcondt{2} 4| rcondf{1} 4;2:rcondf{2} 4]; + 1,2,4,5:(by move=>?;conseq (_:true);auto);2:by sim. + inline *;rcondt{1} 7;1:by auto=>/>. + by wp;rnd;auto;rnd{1};auto => /> *; rewrite !get_setE. + + + proc;sp;if=>//;sim. + call (_: ={FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries});2:by auto. + if=> //;2:sim. + swap{1} 8 -3. + seq 6 6 : (={y1,hx2,t,x,FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries} + /\ (t = in_dom_with FRO.m (oget G1.mhi.[(x.`1, hx2)]).`2 Unknown){1}). + + by inline *;auto. + case (((x.`1, hx2) \in G1.mhi /\ t){1}); + [rcondt{1} 3;2:rcondt{2} 3| rcondf{1} 3;2:rcondf{2} 3]; + 1,2,4,5:(by move=>?;conseq (_:true);auto);2:by sim. + inline *;rcondt{1} 6;1:by auto=>/>. + by auto => /> *; rewrite !get_setE. + + proc;sp;if=>//;auto;if;1:auto;sim. + call (_: ={FRO.m,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.chandle,G1.paths,C.c,C.queries});2:by auto. + by inline*;sim. +qed. + +local module G4(RO:F.RO) = { + + module C = { + + proc f(p : block list): block = { + var sa; + var h, i <- 0; + sa <- b0; + while (i < size p ) { + RO.sample(take (i+1) p); + i <- i + 1; + } + sa <@ RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2; + + if (x \notin G1.m) { + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + y1 <@ RO.get (rcons p (v +^ x.`1)); + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + G1.m.[x] <- y; + G1.mi.[y] <- x; + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + G1.paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget G1.m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2; + + if (x \notin G1.mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + G1.mi.[x] <- y; + G1.m.[y] <- x; + } else { + y <- oget G1.mi.[x]; + } + return y; + } + + } + + proc distinguish(): bool = { + var b; + + RO.init(); + G1.m <- empty; + G1.mi <- empty; + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + G1.paths <- empty.[c0 <- ([<:block>],b0)]; + b <@ DRestr(D,C,S).distinguish(); + return b; + } +}. + +local equiv G3_G4 : G3(F.RO).distinguish ~ G4(F.RO).distinguish : ={glob D} ==> ={res}. +proof. + proc;inline *;wp. + call (_: ={G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});last by auto. + + proc;sp;if=>//;sim. + call (_: ={G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});last by auto. + if => //;2:sim. + seq 3 3: (={x,y1,y2,y,G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});1:by sim. + sim;seq 5 0: (={x,y1,y2,y,G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});1:by inline *;auto. + by if{1};sim;inline *;auto. + + proc;sp;if=>//;sim. + call (_: ={G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});last by auto. + if => //;2:sim. + seq 5 0: (={x,G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});1:by inline *;auto. + seq 3 3: (={x,y1,y2,y,G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});1:by sim. + by if{1};sim;inline *;auto. + proc;sp;if=>//;auto;if=>//;sim. + call (_: ={G1.m,G1.mi,G1.paths,F.RO.m,C.c,C.queries});last by auto. + sp;sim; while(={i,p,F.RO.m})=>//. + inline F.RO.sample F.RO.get;if{1};1:by auto. + if{1};2:by auto. + by sim;inline *;auto;progress;smt(DCapacity.dunifin_ll). +qed. + +local equiv G4_Ideal : G4(F.LRO).distinguish ~ IdealIndif(IF,S,DRestr(D)).main : + ={glob D} ==> ={res}. +proof. + proc;inline *;wp. + call (_: ={C.c,C.queries,F.RO.m} /\ G1.m{1}=S.m{2} /\ G1.mi{1}=S.mi{2} /\ G1.paths{1}=S.paths{2}). + + by sim. + by sim. + + proc;sp;if=>//;auto;if=>//;auto. + call (_: ={F.RO.m});2:by auto. + inline F.LRO.get F.FRO.sample;wp 7 2;sim. + by while{1} (true) (size p - i){1};auto;1:inline*;auto=>/#. + by auto. +qed. + +declare axiom D_ll : + forall (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}), + islossless P.f => + islossless P.fi => islossless F.f => islossless D(F, P).distinguish. + + +lemma Real_Ideal &m: + Pr[GReal(D).main() @ &m: res /\ C.c <= max_size] <= + Pr[IdealIndif(IF,S,DRestr(D)).main() @ &m :res] + + (max_size ^ 2 - max_size)%r / 2%r * mu dstate (pred1 witness) + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). +proof. + apply (ler_trans _ _ _ (Real_G2 D D_ll &m)). + rewrite !(ler_add2l, ler_add2r);apply lerr_eq. + apply (eq_trans _ Pr[G3(F.LRO).distinguish() @ &m : res]);1:by byequiv G2_G3. + apply (eq_trans _ Pr[G3(F.RO ).distinguish() @ &m : res]). + + byequiv (_: ={glob G3, F.RO.m} ==> _)=>//;symmetry;conseq (F.FullEager.RO_LRO_D G3 _)=> //. + by move=> _; exact/Block.DBlock.dunifin_ll. + apply (eq_trans _ Pr[G4(F.RO ).distinguish() @ &m : res]);1:by byequiv G3_G4. + apply (eq_trans _ Pr[G4(F.LRO).distinguish() @ &m : res]). + + byequiv (F.FullEager.RO_LRO_D G4 _)=> //. + by move=> _; exact/Block.DBlock.dunifin_ll. + by byequiv G4_Ideal. +qed. + +end section. diff --git a/sha3/proof/smart_counter/Gconcl_list.ec b/sha3/proof/smart_counter/Gconcl_list.ec new file mode 100644 index 0000000..a9235b8 --- /dev/null +++ b/sha3/proof/smart_counter/Gconcl_list.ec @@ -0,0 +1,2199 @@ +pragma -oldip. +require import Core Int Real StdOrder Ring StdBigop. +require import List FSet SmtMap Common SLCommon PROM FelTactic Mu_mem. +require import Distr DProd Dexcepted BlockSponge Gconcl. +(*...*) import Capacity IntID IntOrder Bigreal RealOrder BRA. + +require (*--*) Handle. + +(*** THEORY PARAMETERS ***) +(** Validity of Functionality Queries **) +op valid: block list -> bool = valid_block. +axiom valid_spec p: valid p => p <> []. +axiom parse_gt0 x: 0 < (parse x).`2. +axiom parse_not_valid x : + !valid (parse x).`1 => + forall i, ! valid (parse (format (parse x).`1 i)).`1. +axiom parse_twice p n x : + (p,n) = parse x => forall i, 0 < i <= n => parse (format p i) = (p,i). +axiom valid_uniq p1 p2 n1 n2 : + valid p1 => valid p2 => format p1 n1 = format p2 n2 => p1 = p2 /\ n1 = n2. + + +clone export Handle as Handle0. + +module DSqueeze (F : SLCommon.DFUNCTIONALITY) = { + proc init () : unit = {} + proc f (p : block list, n : int) : block list = { + var lres : block list <- []; + var b : block <- b0; + var i : int <- 0; + if (valid p) { + b <@ F.f(p); + while (i < n) { + i <- i + 1; + lres <- rcons lres b; + if (i < n) { + b <@ F.f(format p (i+1)); + } + } + } + return lres; + } +}. + + +module (Squeeze (F : SLCommon.FUNCTIONALITY) : FUNCTIONALITY) = { + proc init () : unit = { + C.init(); + F.init(); + } + proc f = DSqueeze(F).f +}. + + +module (A (D : DISTINGUISHER) : SLCommon.DISTINGUISHER) + (F : SLCommon.DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish() : bool = { + var b : bool; + C.init(); + b <@ DRestr(D,DSqueeze(F),P).distinguish(); + return b; + } +}. + + + +module NIndif (F : FUNCTIONALITY, P : PRIMITIVE, D : DISTINGUISHER) = { + proc main () : bool = { + var b : bool; + C.init(); + P.init(); + F.init(); + b <@ D(F,P).distinguish(); + return b; + } +}. + + + +module P = Perm. + +clone IRO as BIRO2 with + type from <- block list, + type to <- block, + op valid <- predT, + op dto <- bdistr. + +module Valid (F : DFUNCTIONALITY) = { + proc init () = {} + proc f (q : block list, k : int) = { + var re : block list <- []; + (q,k) <- parse (format q k); + if (valid q) { + re <@ F.f(q,k); + } else { + re <@ BIRO2.IRO.f(q,k); + } + return re; + } +}. + +module SimLast (S : SLCommon.SIMULATOR) (F : DFUNCTIONALITY) = { + proc init() = { + BIRO2.IRO.init(); + S(Last(Valid(F))).init(); + } + proc f = S(Last(Valid(F))).f + proc fi = S(Last(Valid(F))).fi +}. + +clone F as F2. + + +section Ideal. + + op (<=) (m1 m2 : (block list, 'b) fmap) = + forall x, x <> [] => x \in m1 => m1.[x] = m2.[x]. + + local lemma leq_add_nin (m1 m2 : (block list, 'b) fmap) (x : block list) (y : 'b): + m1 <= m2 => + ! x \in m2 => + m1 <= m2.[x <- y]. + proof. + move=>h_leq H_n_dom a H_a_dom;rewrite get_setE/=;smt(domE). + qed. + + + local lemma leq_add_in (m1 m2 : (block list, 'b) fmap) (x : block list) : + m1 <= m2 => + x \in m2 => + m1.[x <- oget m2.[x]] <= m2. + proof. + move=>h_leq H_n_dom a H_a_dom;rewrite get_setE/=;smt(domE get_setE). + qed. + + local lemma leq_nin_dom (m1 m2 : (block list, 'b) fmap) (x : block list) : + m1 <= m2 => + x <> [] => + ! x \in m2 => ! x \in m1 by smt(domE). + + local lemma prefix_leq1 (l : block list) (m : (block list,block) fmap) i : + 0 <= i => + format l (i+1) \in m => + size (format l (i+1)) <= prefix (format l (i+1+1)) + (get_max_prefix (format l (i+1+1)) (elems (fdom m))) <= size (format l (i+1+1)). + proof. + rewrite -mem_fdom memE;move=>hi0 H_dom. + have->:(format l (i + 1 + 1)) = format l (i + 1) ++ [b0]. + + by rewrite/format//=nseqSr//-cats1 catA. + have:=prefix_leq_prefix_cat_size (format l (i + 1))[b0](elems (fdom m)). + rewrite (prefix_get_max_prefix_eq_size _ _ H_dom)//=. + rewrite (size_cat _ [b0])/=;pose x:= format _ _. + have:=get_max_prefix_max (x ++ [b0]) _ _ H_dom. + have->:prefix (x ++ [b0]) (format l (i + 1)) = size x + by rewrite prefixC-{1}(cats0 (format l (i+1)))/x prefix_cat//=. + smt(prefix_sizel size_cat prefix_ge0 ). + qed. + + local lemma prefix_le1 (l : block list) (m : (block list,block) fmap) i : + 0 <= i => + format l (i+1) \in m => + size (format l (i+1+1)) - prefix (format l (i+1+1)) + (get_max_prefix (format l (i+1+1)) (elems (fdom m))) <= 1. + proof. + move=> Hi0 H_liS_in_m. + have:= prefix_leq1 _ _ _ Hi0 H_liS_in_m. + rewrite /format /= nseqSr //-cats1 catA (size_cat(l ++ nseq i b0) [b0]) /=. + smt(size_ge0). + qed. + + local lemma leq_add2 (m1 m2 : (block list, 'b) fmap) (x : block list) (y : 'b) : + m1 <= m2 => + ! x \in m2 => + m1.[x <- y] <= m2.[x <- y] by smt(domE get_setE mem_set in_fsetU1). + + + local equiv ideal_equiv (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S}) : + SLCommon.IdealIndif(IF, S, SLCommon.DRestr(A(D))).main + ~ + SLCommon.IdealIndif(IF, S, A(D)).main + : + ={glob D} ==> ={glob D, res}. + proof. + proc;inline*;auto;sp. + call(: ={glob IF, glob S, glob A} /\ SLCommon.C.c{1} <= C.c{1} + /\ SLCommon.C.queries{1} <= F.RO.m{2});auto;last first. + + progress. + by move=>x;rewrite get_setE/=mem_set-mem_fdom fdom0 in_fset0//==>->. + + proc;inline*;sp;if;auto;sp;rcondt{1}1;1:auto=>/#;sp;if=>//=;2:auto=>/#. + wp 7 6;conseq(:_==> ={y} /\ ={F.RO.m} /\ ={S.paths, S.mi, S.m} + /\ SLCommon.C.queries{1} <= F.RO.m{2});1:smt(). + if;auto;smt(leq_add_nin). + + by proc;inline*;sp;if;auto;sp;rcondt{1}1;1:auto=>/#;sp;if;auto;smt(). + proc;inline*;sp;if;auto;swap 6;auto;sp;if;auto;2:smt(size_ge0). + case(0 < n{1});last first. + + sp;rcondf{1}3;2:rcondf{2}4;1,2:auto. + - by if;auto;if;auto. + by if{1};2:auto;1:if{1};auto; + smt(prefix_ge0 leq_add_in DBlock.dunifin_ll domE size_ge0 get_setE leq_add2). + splitwhile{1}5: i + 1 < n;splitwhile{2}5: i + 1 < n. + rcondt{1}6;2:rcondt{2}6;auto. + * by while(i < n);auto;sp;if;auto;sp;if;auto;if;auto. + * by while(i < n);auto;sp;if;auto;sp;if;auto;if;auto. + rcondf{1}8;2:rcondf{2}8;auto. + * while(i < n);auto. + by sp;if;auto;sp;if;auto;if;auto. + sp;if;auto;2:smt();if;auto;smt(). + * while(i < n);auto;2:smt();sp;if;auto;sp;if;auto;if;auto. + rcondf{1}8;2:rcondf{2}8;auto. + * while(i < n);auto. + by sp;if;auto;sp;if;auto;if;auto. + sp;if;auto;2:smt();if;auto;smt(). + * by while(i < n);auto;2:smt();sp;if;auto;sp;if;auto;if;auto. + conseq(:_==> ={b,lres,F.RO.m,S.paths,S.mi,S.m} + /\ i{1} = n{1} - 1 + /\ SLCommon.C.c{1} <= C.c{1} + size bl{1} + i{1} + /\ SLCommon.C.queries{1} <= F.RO.m{2});1:smt(). + while(={lres,F.RO.m,S.paths,S.mi,S.m,i,n,p,nb,b,bl} + /\ 0 <= i{1} <= n{1} - 1 + /\ SLCommon.C.queries.[format p (i+1)]{1} = Some b{1} + /\ p{1} = bs{1} /\ valid p{1} /\ p{1} = bl{1} + /\ C.c{1} + size p{1} + n{1} - 1 <= max_size + /\ SLCommon.C.c{1} <= C.c{1} + size bl{1} + i{1} + /\ SLCommon.C.queries{1} <= F.RO.m{2});progress. + sp;rcondt{1}1;2:rcondt{2}1;1,2:auto;sp. + case((x0 \in F.RO.m){2});last first. + * rcondt{2}2;1:auto;rcondt{1}1. + + auto => /> &hr iR 9?; apply leq_nin_dom => //. + smt (leq_nin_dom size_cat size_eq0 size_nseq valid_spec). + rcondt{1}1;1:auto. + - move=> /> &hr i [#] h1 h2 h3 h4 h5 h6 h7 h8 h9 h10. + have//= /#:= prefix_le1 bl{m} SLCommon.C.queries{hr} i h1 _. + by rewrite domE h3. + sp;rcondt{1}2;auto;progress. + - smt(). + - smt(). + - by rewrite!get_setE/=. + - have//= /#:= prefix_le1 bl{2} SLCommon.C.queries{1} i_R H _. + by rewrite domE H1. + - by rewrite!get_setE/= leq_add2//=. + if{1}. + * rcondt{1}1;1:auto. + - move=> /> &hr i [#] h1 h2 h3 h4 h5 h6 h7 h8 h9 h10. + have//= /#:= prefix_le1 bl{m} SLCommon.C.queries{hr} i h1 _. + by rewrite domE h3. + sp;rcondf{1}2;2:rcondf{2}2;auto;progress. + - smt(). + - smt(). + - by rewrite!get_setE/=. + - have//= /#:= prefix_le1 bl{2} SLCommon.C.queries{1} i_R H _. + by rewrite domE H1. + - smt(leq_add_in domE). + rcondf{2}2;auto;progress. + - smt(size_cat size_nseq size_eq0 size_ge0). + - smt(). + - smt(). + - by move: H11; rewrite domE; case: (SLCommon.C.queries{1}.[format bl{2} (i_R + 2)]). + - smt(). + sp;conseq(:_==> ={F.RO.m,b} + /\ SLCommon.C.queries.[p]{1} = Some b{1} + /\ SLCommon.C.c{1} <= C.c{1} + size bl{1} + /\ SLCommon.C.queries{1} <= F.RO.m{2});progress. + - smt(). + - smt(nseq0 cats0). + - smt(size_ge0). + - smt(). + case(p{2} \in F.RO.m{2}). + + rcondf{2}2;1:auto. + sp;if{1}. + - rcondt{1}1;1:auto;1:smt(prefix_ge0). + sp;rcondf{1}2;auto;progress. + * by rewrite!get_setE/=. + * smt(prefix_ge0). + * smt(leq_add_in domE). + auto;progress. + - smt(domE). + - smt(domE). + - smt(size_ge0). + rcondt{1}1;1:auto;1:smt(leq_nin_dom domE). + rcondt{1}1;1:auto;1:smt(prefix_ge0). + sp;auto;progress. + + by rewrite!get_setE/=. + + smt(prefix_ge0). + + rewrite get_setE/= leq_add2//=. + + by rewrite!get_setE/=. + + smt(prefix_ge0). + + exact leq_add_in. + qed. + + + local module IF'(F : F.RO) = { + proc init = F.init + proc f (x : block list) : block = { + var b : block <- b0; + var i : int <- 0; + var p,n; + (p,n) <- parse x; + while (i < n) { + i <- i + 1; + F.sample(format p i); + } + b <@ F.get(x); + return b; + } + }. + + + + + local module (L (D : DISTINGUISHER) : F.RO_Distinguisher) (F : F.RO) = { + proc distinguish = SLCommon.IdealIndif(IF'(F), S, A(D)).main + }. + + local module Valid2 (F : F.RO) = { + proc init = F.init + proc f (q : block list) = { + var r : block <- b0; + var s,t; + (s,t) <- parse q; + r <@ F.get(q); + return r; + } + }. + + local module (L2 (D : DISTINGUISHER) : F.RO_Distinguisher) (F : F.RO) = { + proc distinguish = SLCommon.IdealIndif(Valid2(F), S, A(D)).main + }. + + local equiv Ideal_equiv_valid (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S}) : + L(D,F.LRO).distinguish + ~ + L2(D,F.LRO).distinguish + : + ={glob D} ==> ={glob D, res}. + proof. + proc;inline*;sp;wp. + call(: ={glob S, glob C, glob F.RO});auto. + + proc;sp;if;auto. + call(: ={glob S,glob F.RO});auto. + sp;if;auto;if;auto;sp. + call(: ={glob F.RO});2:auto;2:smt(). + inline F.LRO.sample;call(: ={glob IF});auto;progress. + by while{1}(true)(n{1}-i{1});auto;smt(). + + by proc;sim. + proc;sp;if;auto;sp;call(: ={glob IF,glob S});auto. + sp;if;auto. + while(={glob S,glob IF,lres,i,n,p,b}). + + sp;if;auto. + call(: ={glob IF});auto. + call(: ={glob IF});auto. + conseq(:_==> true);auto. + by inline*;while{1}(true)(n{1}-i{1});auto;smt(). + call(: ={glob IF});auto. + call(: ={glob IF});auto. + conseq(:_==> true);auto. + by inline*;while{1}(true)(n{1}-i{1});auto;smt(). + qed. + + + local equiv ideal_equiv2 (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S}) : + L2(D,F.RO).distinguish ~ SLCommon.IdealIndif(IF,S,A(D)).main + : ={glob D} ==> ={glob D, res}. + proof. + proc;inline*;sp;wp. + call(: ={glob F.RO, glob S, glob C});auto. + + proc;auto;sp;if;auto. + call(: ={glob F.RO, glob S});auto. + if;1,3:auto;sim;if;auto. + call(: ={glob F.RO});2:auto. + by inline*;sp;wp 2 2;sim. + + by proc;sim. + proc;sp;if;auto;sp. + call(: ={glob F.RO});auto;sp;if;auto;inline*;auto;sp. + case(0 < n{1});last first. + + by rcondf{2}4;1:auto;rcondf{1}5;auto. + while(={lres,F.RO.m,i,n,p,b} /\ valid p{1} /\ 0 <= i{1} <= n{1}). + + sp;if;1:auto. + - by auto;smt(parse_valid parseK formatK). + by auto;smt(parse_valid parseK formatK). + by auto;smt(parse_valid parseK formatK). + qed. + + inductive inv_L_L3 (m1 m2 m3 : (block list, block) fmap) = + | INV of (m1 = m2 + m3) + & (forall l, l \in m2 => valid (parse l).`1) + & (forall l, l \in m3 => ! valid (parse l).`1). + + local module IF2(F : F.RO) (F2 : F2.RO) = { + proc init () = { + F.init(); + F2.init(); + } + proc f (x : block list) : block = { + var b : block <- b0; + var i : int <- 0; + var p,n; + (p,n) <- parse x; + if (valid p) { + while (i < n) { + i <- i + 1; + F.sample(format p i); + } + b <@ F.get(x); + } else { + while (i < n) { + i <- i + 1; + F2.sample(format p i); + } + b <@ F2.get(x); + } + return b; + } + }. + + + local module (L3 (D : DISTINGUISHER) : F.RO_Distinguisher) (F : F.RO) = { + proc distinguish = SLCommon.IdealIndif(IF2(F,F2.RO), S, A(D)).main + }. + + local lemma lemma1 m1 m2 m3 p i r: + inv_L_L3 m1 m2 m3 => + valid p => + 0 < i => + ! format p i \in m1 => + ! format p i \in m2 => + inv_L_L3 m1.[format p i <- r] m2.[format p i <- r] m3. + proof. + move=>INV0 p_valid i_gt0 nin_dom1 nin_dom2;split;have[]add_maps valid_dom nvalid_dom:=INV0. + + rewrite add_maps -fmap_eqP=>x. + by rewrite get_setE !joinE get_setE;smt(parseK formatK). + + smt(mem_set parseK formatK). + + smt(mem_set parseK formatK). + qed. + + local lemma lemma2 m1 m2 m3 p i: + inv_L_L3 m1 m2 m3 => + valid p => + 0 < i => + format p i \in m1 => + format p i \in m2. + proof. + move=>INV0 p_valid i_gt0 domE1;have[]add_maps valid_dom nvalid_dom:=INV0. + by have:= domE1; rewrite add_maps mem_join;smt(parseK formatK). + qed. + + + local lemma incl_dom m1 m2 m3 l : + inv_L_L3 m1 m2 m3 => + l \in m1 <=> (l \in m2 \/ l \in m3). + proof. + move=>INV0;have[]add_maps valid_dom nvalid_dom:=INV0. + by rewrite add_maps mem_join. + qed. + + + local lemma lemma3 m1 m2 m3 x r: + inv_L_L3 m1 m2 m3 => + ! valid (parse x).`1 => + ! x \in m1 => + inv_L_L3 m1.[x <- r] m2 m3.[x <- r]. + proof. + move=>INV0 not_valid nin_dom1;have[]add_maps h1 h2:=INV0. + have nin_dom3: ! x \in m3 by smt(incl_dom). + split. + + by apply/fmap_eqP=>y;rewrite add_maps !get_setE!joinE!get_setE mem_set/#. + + exact h1. + smt(mem_set). + qed. + + + local equiv Ideal_equiv3 (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S, -F2.RO}) : + L(D,F.RO).distinguish ~ L3(D,F.RO).distinguish + : ={glob D} ==> ={glob D, res}. + proof. + proc;inline*;auto;sp. + call(: ={glob S, glob C} /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});auto;first last. + + proc;sp;if;auto. + by call(: ={glob S});auto;sim. + + proc;sp;if;auto;call(: inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});auto;sp. + inline*;if;auto;sp. + rcondt{1}1;1:auto;1:smt(parse_valid parseK formatK). + rcondt{2}1;1:auto;1:smt(parse_valid parseK formatK). + rcondt{2}1;1:auto;1:smt(parse_valid parseK formatK);sp. + rcondf{1}3;1:auto;1:smt(parse_valid parseK formatK);sp. + rcondf{2}3;1:auto;1:smt(parse_valid parseK formatK);sp. + rcondf{1}5;2:rcondf{2}5; + 1,2:by auto;smt(mem_set nseq0 cats0 parse_valid). + case(0 < n{1});auto;last first. + - rcondf{1}7;1:auto;rcondf{2}7;1:auto. + by wp;rnd;auto;progress;smt(lemma1 nseq0 cats0 lemma2 incl_dom + parse_valid parseK formatK in_fsetU). + while(={i,n,p,lres,b} /\ valid p{1} /\ 0 <= i{1} <= n{1} + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2}). + - sp;if;1,3:auto=>/#. + sp;rcondt{2}1;1:auto;1:smt(parse_valid parseK formatK). + conseq(:_==> ={b} /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});1:progress=>/#. + auto=>/=. + conseq(:_==> inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});progress. + * by rewrite!get_setE//=. + * smt(lemma1 parse_valid). + * smt(lemma2 parse_valid). + * smt(lemma2 parse_valid). + * smt(incl_dom). + * smt(incl_dom). + * case:H8;smt(joinE). + while(={i1,n1,p1} /\ valid p1{1} /\ 0 <= i1{1} <= n1{1} + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2}). + * sp;conseq(:_==> inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});1:smt(). + case(x6{1} \in F.RO.m{1}). + + by rcondf{1}2;2:rcondf{2}2;auto;smt(incl_dom lemma2). + by rcondt{1}2;2:rcondt{2}2;auto;smt(lemma2 incl_dom lemma1). + by auto;smt(parseK). + wp;rnd;wp 2 2. + conseq(:_==> F.RO.m{1}.[p{1}] = F.RO.m{2}.[p{2}] + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});progress. + + have[]add_maps h1 h2:=H5;rewrite add_maps joinE//=. + by have:= h2 p{2}; rewrite parse_valid //= H2 /= => h; rewrite h. + + smt(). + case(x5{1} \in F.RO.m{1}). + - rcondf{1}2;2:rcondf{2}2;auto;progress. + * smt(lemma2 incl_dom parse_valid). + by have[]add_maps h1 h2:=H1;rewrite add_maps joinE//=;smt(parse_valid). + rcondt{1}2;2:rcondt{2}2;auto;progress. + - move:H4;rewrite/format/=nseq0 !cats0 => p0_notin_ROm_m. + case: H1 => joint _ _; move: p0_notin_ROm_m. + by rewrite joint mem_join negb_or; smt(parse_valid). + - have[]add_maps h1 h2:=H1;rewrite add_maps !get_setE joinE//=;smt(parse_valid nseq0 cats0). + - have:=H;rewrite -H0=>//=[][]->>->>;apply lemma1=>//=;1:smt(parse_valid). + have[]add_maps h1 h2:=H1;smt(parse_valid formatK parseK incl_dom). + + progress;split. + - by apply/fmap_eqP=>x;rewrite joinE mem_empty. + - smt(mem_empty). + - smt(mem_empty). + proc;sp;if;auto;call(: ={glob S} /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});auto. + if;1,3:auto. + seq 1 1 : (={x, y1, S.paths, S.mi, S.m} /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});last first. + + by conseq(:_==> ={y, S.paths, S.mi, S.m});progress;sim. + if;auto. + call(: inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});auto;sp;inline*. + if{2}. + + seq 1 1 : (={x,p,n} /\ parse x{1} = (p,n){1} /\ valid p{1} + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});last first. + - sp;case(x1{1} \in F.RO.m{1}). + * rcondf{1}2;2:rcondf{2}2;auto;progress. + + have:=H2;rewrite -formatK H/=;smt(lemma2 incl_dom parse_gt0). + have[]add_maps h1 h2:=H1;rewrite add_maps joinE. + have:=H2;rewrite -formatK H/==>in_dom1. + case(format p{2} n{2} \in F2.RO.m{2})=>//=in_dom3. + by have:=h2 _ in_dom3;rewrite parseK//=;smt(parse_gt0). + rcondt{1}2;2:rcondt{2}2;auto;progress. + + smt(incl_dom lemma2). + + have[]:=H1;smt(get_setE joinE). + by have:=H2;rewrite-formatK H/==>nin_dom1;rewrite lemma1//=;smt(parse_gt0 lemma2 incl_dom). + conseq(:_==> inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});1:smt(). + while(={i,n,p} /\ 0 <= i{1} /\ valid p{1} /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2}). + + sp;case(x2{1} \in F.RO.m{1}). + - by rcondf{1}2;2:rcondf{2}2;auto;smt(lemma2). + rcondt{1}2;2:rcondt{2}2;auto;progress. + - smt(incl_dom lemma1). + - smt(incl_dom lemma1). + apply/lemma1=> //=. + - smt(). + smt(incl_dom mem_join). + auto;smt(). + seq 1 1 : (={x,p,n} /\ parse x{1} = (p,n){1} /\ ! valid p{1} + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});last first. + + sp;case(x1{1} \in F.RO.m{1}). + - rcondf{1}2;2:rcondf{2}2;auto;progress. + * have[]:=H1;smt(incl_dom). + have[]:=H1;smt(joinE incl_dom). + rcondt{1}2;2:rcondt{2}2;auto;progress. + * have[]:=H1;smt(incl_dom). + * have[]:=H1;smt(joinE incl_dom get_setE). + by rewrite(lemma3 _ _ _ _ rL H1 _ H2)H//=. + conseq(:_==> inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2});1:smt(). + while(={i,n,p,x} /\ 0 <= i{1} /\ ! valid p{1} /\ parse x{1} = (p,n){1} + /\ inv_L_L3 F.RO.m{1} F.RO.m{2} F2.RO.m{2}). + + sp;case(x2{1} \in F.RO.m{1}). + - rcondf{1}2;2:rcondf{2}2;auto;progress. + * have[]h_join h1 h2:=H2. + have:= H5; rewrite h_join mem_join. + have:= h1 (format p{hr} (i_R + 1)). + have:=parse_not_valid x{hr}; rewrite H1 /= H0 /= => h. + by rewrite (h (i_R+1)) /= => ->. + smt(). + rcondt{1}2;2:rcondt{2}2;auto;progress. + * smt(incl_dom lemma1). + * smt(). + * have//=:=lemma3 _ _ _ _ r0L H2 _ H5. + by have:= parse_not_valid x{2}; rewrite H1 /= H0 /= => h; exact/(h (i_R+1)). + auto;smt(). + qed. + + local module D2 (D : DISTINGUISHER) (F : F.RO) = { + proc distinguish = D(FC(DSqueeze(Valid2(F))), PC(S(Valid2(F)))).distinguish + }. + + local module D3 (D : DISTINGUISHER) (F : F.RO) = { + proc distinguish = D(FC(DSqueeze(IF'(F))), PC(S(IF'(F)))).distinguish + }. + + + module DSqueeze2 (F : F.RO) (F2 : F2.RO) = { + proc init () : unit = { + F.init(); + F2.init(); + } + proc f (p : block list, n : int) : block list = { + var lres : block list <- []; + var b : block <- b0; + var i : int <- 0; + var pp, nn; + (pp,nn) <- parse (format p n); + if (valid p) { + if (n <= 0) { + F.sample(p); + } + while (i < n) { + i <- i + 1; + b <@ F.get(format p i); + lres <- rcons lres b; + } + } else { + if (nn <= 0) { + F2.sample(pp); + } + while (i < nn - n) { + i <- i + 1; + F2.sample(format pp i); + } + while (i < n) { + i <- i + 1; + b <@ F2.get(format pp i); + lres <- rcons lres b; + } + } + return lres; + } + }. + + + local module FValid (F : DFUNCTIONALITY) = { + proc f (p : block list, n : int) = { + var r : block list <- []; + if (valid p) { + r <@ F.f(p,n); + } + return r; + } + }. + + local module DValid (D : DISTINGUISHER) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = D(FValid(F),P). + + local module S2 (F : DFUNCTIONALITY) = S(Last(F)). + + local module L4 (D : DISTINGUISHER) (F : F.RO) (F2 : F2.RO) = { + proc distinguish = IdealIndif(DSqueeze2(F,F2),S2,DValid(DRestr(D))).main + }. + + local equiv equiv_L3_L4 (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S, -F2.RO, -BIRO.IRO, -BIRO2.IRO}) : + L3(D,F.RO).distinguish + ~ + L4(D,F.RO,F2.RO).distinguish + : + ={glob D} ==> ={glob D, res}. + proof. + proc; inline*; auto; sp. + call(: ={glob S, glob C, glob F.RO, glob F2.RO}); auto;first last. + + by proc; sim. + + proc; sp; if; auto; call(: ={glob F.RO, glob F2.RO}); auto; sp; if; auto; inline*; sp. + rcondt{1}1; 1:(auto; smt(parse_valid parse_gt0)); sp. + rcondt{1}1; 1:(auto; smt(parse_valid parse_gt0)); sp. + (* rcondt{1}1; 1:(auto; smt(parse_valid parse_gt0)); sp. *) + rcondf{1}3; 1:(auto; smt(parse_valid parse_gt0)); sp. + rcondt{2}1; 1:(auto; smt(parse_valid parse_gt0 parseK formatK)); sp; wp. + if{2};sp. + - rcondf{2}3; 1:(auto; smt(parse_valid parse_gt0)); sp. + rcondf{1}8; 1:(auto; smt(parse_valid parse_gt0)); sp. + rcondf{1}5; 1:(auto; smt(parse_valid parse_gt0 mem_set nseq0 cats0)); sp. + wp 4 2;rnd{1};wp 2 2. + by conseq(:_==> ={F.RO.m} /\ r3{1} = r2{2} /\ x9{1} = x4{2});2:sim; + smt(DBlock.dunifin_ll nseq0 cats0 parse_valid);progress. + rcondt{2}1; 1:(auto; smt(parse_valid parse_gt0)); sp; wp. + splitwhile{1} 8 : i + 1 < n. + rcondt{1}9;1:auto. + - by while(i < n);auto;2:smt();sp;if;auto;1:(sp;if;auto);while(i < n);auto. + rcondf{1}11;1:auto. + - by while(i < n);auto;2:smt();sp;if;auto;1:(sp;if;auto);while(i < n);auto. + rcondf{1}11;1:auto. + - by while(i < n);auto;2:smt();sp;if;auto;1:(sp;if;auto);while(i < n);auto. + wp. + while((n,p){1} = (n0,p0){2} /\ i{1} + 1 = i{2} /\ valid p{1} /\ 0 < n{1} + /\ 0 <= i{2} <= n{1} + /\ (forall j, 1 <= j <= i{2} => format p{1} j \in F.RO.m{1}) + /\ rcons lres{1} b{1} = lres{2} /\ ={F.RO.m, F2.RO.m});last first. + - rcondf{1}5;1:auto;1:smt(mem_set nseq0 cats0 parse_valid). + wp 4 2;rnd{1};wp 2 2. + conseq(:_==> ={F.RO.m} /\ r3{1} = r0{2} /\ x9{1} \in F.RO.m{1}); + 1:smt(DBlock.dunifin_ll nseq0 cats0 parse_valid). + by auto;smt(parse_valid nseq0 cats0 mem_set). + sp. + rcondt{1}1;1:auto;sp. + rcondt{1}1;1:(auto;smt(parse_valid parseK formatK)). + (* rcondt{1}1;1:(auto;smt(parse_valid parseK formatK parse_gt0)). *) + splitwhile{1} 1 : i1 + 1 < n1. + rcondt{1}2;1:auto. + - by while(i1 < n1);auto;smt(parse_gt0 parse_valid parseK formatK). + rcondf{1}7;1:auto. + - by while(i1 < n1);auto;smt(parse_gt0 parse_valid parseK formatK). + rcondf{1}9;1:auto. + - conseq(:_==> i1 + 1 = n1);1:smt(mem_set parseK parse_valid formatK). + by while(i1 + 1 <= n1);auto;smt(parse_gt0 parse_valid parseK formatK). + wp 8 2;rnd{1};wp 6 2. + conseq(:_==> n1{1} = i{2} /\ ={F.RO.m} /\ i1{1} = n1{1} + /\ (forall (j : int), 1 <= j <= i{2} => + format p1{1} j \in F.RO.m{1})); + 1:smt(parseK formatK parse_valid DBlock.dunifin_ll). + seq 2 0 : (={F.RO.m,x0} /\ i1{1} = n1{1} /\ x0{2} = format p{1} i{2} + /\ n1{1} = i{1} + 1 /\ p1{1} = p{1} /\ i{2} = i{1} + 1 /\ forall (j : int), + 1 <= j <= i{1} => format p{1} j \in F.RO.m{1});last first. + - auto;smt(mem_set). + wp;conseq(:_==> ={F.RO.m} /\ i1{1} + 1 = n1{1} + /\ (forall (j : int), 1 <= j < n1{1} => + format p1{1} j \in F.RO.m{1}));1:smt(parseK). + while{1}(={F.RO.m} /\ 0 <= i1{1} /\ i1{1} + 1 <= n1{1} /\ i{2} = n1{1} /\ i{2} = i{1} + 1 + /\ (forall (j : int), 1 <= j < n1{1} => + format p1{1} j \in F.RO.m{1}))(n1{1}-i1{1}-1);progress. + + by sp;rcondf 2;auto;smt(DBlock.dunifin_ll). + by auto;smt(parse_gt0 parseK formatK parse_valid). + proc; sp; if; auto; call(: ={glob S, glob F.RO, glob F2.RO}); auto. + if; 1,3:auto; sim; if; auto; sim; sp. + call(: ={glob F.RO, glob F2.RO});auto;last smt(). + inline*;auto;sp. + if;1:auto;1:smt(). + + (* rcondt{1}1;1:(auto;smt(parse_valid parse_gt0)). *) + rcondf{2}1;1:(auto;smt(parse_valid parse_gt0)). + splitwhile{1} 1 : i + 1 < n;splitwhile{2} 1 : i + 1 < n. + rcondt{1}2;1:auto. + - by while(i ={F.RO.m} /\ p{2} = x0{2});progress. + + smt(last_rcons formatK parseK). + seq 3 3 : (={F.RO.m,i,x0} /\ x0{1} = p{2}); + last by conseq(:_==> ={F.RO.m});progress;sim. + auto;conseq(:_==> ={F.RO.m,i,n} /\ p{1} = p0{2} /\ i{1} + 1 = n{2});1:smt(formatK). + by while(={F.RO.m,i,n} /\ p{1} = p0{2} /\ 0 <= i{1} /\ i{1} + 1 <= n{2}); + auto;smt(parse_gt0). + sp;rcondf{2}1;1:(auto;smt(parse_gt0)). + rcondf{2}1;1:auto;1:smt(parseK formatK). + splitwhile{1} 1 : i + 1 < n;splitwhile{2} 1 : i + 1 < n. + rcondt{1}2;1:auto. + - by while(i ={F2.RO.m} /\ format pp{2} n{2} = x3{2}). + + move=> /> &1 &2 H H0 /= /> [#] H1 H2 m lres /= ?. + smt(last_rcons formatK parseK). + seq 3 3 : (={F2.RO.m,i} /\ x2{1} = x3{2} /\ pp{2} = p{1} /\ format pp{2} n{2} = x3{2}); + last by conseq(:_==> ={F2.RO.m});progress;sim. + auto;conseq(:_==> ={F2.RO.m,i,n} /\ i{1} + 1 = n{2});1:smt(formatK). + by while(={F2.RO.m,i,n} /\ p{1} = pp{2} /\ 0 <= i{1} /\ i{1} + 1 <= n{2}); + auto;smt(parse_gt0 parseK formatK). + qed. + + + op inv_map (m1 : (block list, block) fmap) (m2 : (block list * int, block) fmap) = + (forall p n x, parse x = (p,n+1) => (p,n) \in m2 <=> x \in m1) + /\ (forall p n x, parse x = (p,n+1) => x \in m1 <=> (p,n) \in m2) + /\ (forall p n x, parse x = (p,n+1) => m2.[(p,n)] = m1.[x]) + /\ (forall p n x, parse x = (p,n+1) => m1.[x] = m2.[(p,n)]). + + inductive INV_L4_ideal m1 m2 m3 m4 = + | inv_maps of (inv_map m1 m2) + & (inv_map m3 m4) + & (forall p n, (p,n) \in m2 => valid p /\ 0 <= n) + & (forall p n, (p,n) \in m4 => ! valid p /\ 0 <= n). + + + local lemma lemma5 m1 m2 m3 m4 p i r : + INV_L4_ideal m1 m2 m3 m4 => + ! (p,i) \in m2 => + 0 <= i => + valid p => + INV_L4_ideal m1.[format p (i+1) <- r] m2.[(p,i) <- r] m3 m4. + proof. + move=>INV0 nin_dom1 i_gt0 valid_p;have[]inv12 inv34 dom2 dom4:=INV0;have[]h1[]h2[]h3 h4:=inv12;split=>//=. + + progress. + - move:H0;rewrite 2!mem_set=>[][/#|]/=[]->>->>;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - smt(get_setE parseK formatK). + smt(get_setE parseK formatK). + smt(get_setE parseK formatK mem_set). + qed. + + local lemma lemma5bis m1 m2 m3 m4 p i r : + INV_L4_ideal m1 m2 m3 m4 => + ! (p,i) \in m4 => + 0 <= i => + ! valid p => + parse (format p (i+1)) = (p,i+1) => + INV_L4_ideal m1 m2 m3.[format p (i+1) <- r] m4.[(p,i) <- r]. + proof. + move=>INV0 nin_dom1 i_gt0 nvalid_p parseK_p_i; + have[]inv12 inv34 dom2 dom4:=INV0; + have[]h1[]h2[]h3 h4:=inv34; + split=>//=. + + progress. + - move:H0;rewrite 2!mem_set=>[][/#|]/=[]->>->>;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - move:H0;rewrite 2!mem_set=>[][/#|]/=;smt(parseK formatK). + - smt(get_setE parseK formatK). + smt(get_setE parseK formatK). + smt(get_setE parseK formatK mem_set). + qed. + + + + local equiv equiv_L4_ideal (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S, -F2.RO, -BIRO.IRO, -BIRO2.IRO}) : + L4(D,F.LRO,F2.LRO).distinguish + ~ + IdealIndif(BIRO.IRO,SimLast(S),DRestr(D)).main + : + ={glob D} ==> ={glob D, res}. + proof. + proc; inline*; auto; sp. + call(: ={glob S, glob C} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2}); + auto; -1:(progress;split;smt(mem_empty in_fset0 emptyE)). + + proc;sp;if;auto;call(: ={glob S} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2}); auto. + if;1,3:auto. seq 1 1 : (={y1, x, glob S} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2}); + last by conseq(:_==> ={y, glob S});progress;sim. + if;auto;call(: INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2});auto. + inline*;auto;sp;if;auto;1:smt(parseK parse_gt0 formatK);1:sp 0 4;2:sp 0 3. + - rcondt{2}1;1:auto;1:smt(parseK parse_gt0 formatK). + while(lres{1} = bs{2} /\ ={i,n} /\ x{2} = p0{1} /\ valid p0{1} /\ 0 <= i{1} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2});progress. + * sp;if{2}. + + rcondt{1}2;auto;progress. + - have[]h1 _ _ _:=H1;have[]h'1 _:=h1;smt(parseK). + - smt(get_setE). + - smt(). + - exact lemma5. + rcondf{1}2;auto;progress. + - have[]h1 _ _ _:=H1;have[]h'1 _:=h1;smt(parseK). + - have[]h1:=H1;have[]:=h1;smt(parseK). + smt(). + by if{1};auto;smt(parseK parse_gt0 formatK). + rcondf{1}1;1:auto;1:smt(parse_gt0);sp. + rcondt{2}1;1:auto. + while(lres{1} = bs0{2} /\ (i,n,pp){1} = (i0,n0,x0){2} + /\ (x0{2}, n0{2}) = parse (format q{2} k{2}) /\ ! valid x0{2} /\ 0 <= i{1} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2});progress. + * sp;if{2}. + + rcondt{1}2;auto;progress. + - have[]_ h1 _ _:=H2;have[]:=h1;progress. + have:=H7 x0{m} i0{m} (format x0{m} (i0{m} + 1));rewrite H5/==>->//=. + have->/#:=parse_twice _ _ _ H. + - smt(get_setE). + - smt(). + - apply lemma5bis=>//=. + rewrite(parse_twice _ _ _ H)/#. + rcondf{1}2;auto;progress. + - have[]_ h1 _ _:=H2;have[]:=h1;progress. + have:=H7 x0{m} i0{m} (format x0{m} (i0{m} + 1));rewrite H5/==>->//=. + have->/#:=parse_twice _ _ _ H. + - have[]_ h1 _ _:=H2;have[]h'1 _:=h1;smt(parseK parse_twice). + - smt(). + by rcondf{1}1;auto;smt(parseK formatK). + + by proc;inline*;conseq(:_==> ={glob C, glob S, z});progress;sim. + proc;sp;if;auto;call(: INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} + F2.RO.m{1} BIRO2.IRO.mp{2});auto. + inline*;sp;if;auto;sp. + rcondt{1}1;1:auto;if{1};sp. + - by rcondf{1}1;2:rcondf{2}1;auto;smt(). + while(lres{1} = bs{2} /\ ={i} /\ n0{1} = n{2} /\ x{2} = p0{1} /\ valid p0{1} /\ 0 <= i{1} + /\ INV_L4_ideal F.RO.m{1} BIRO.IRO.mp{2} F2.RO.m{1} BIRO2.IRO.mp{2});progress. + sp;if{2}. + + rcondt{1}2;auto;progress. + - have[]h1 _ _ _:=H1;have[]h'1 _:=h1;smt(parseK). + - smt(get_setE). + - smt(). + - exact lemma5. + rcondf{1}2;auto;progress. + - have[]h1 _ _ _:=H1;have[]h'1 _:=h1;smt(parseK). + - have[]h1:=H1;have[]:=h1;smt(parseK). + smt(). + qed. + + local module D5 (D : DISTINGUISHER) (F : F.RO) = + D(FC(FValid(DSqueeze2(F, F2.RO))), PC(S(Last(DSqueeze2(F, F2.RO))))). + + local module D6 (D : DISTINGUISHER) (F2 : F2.RO) = + D(FC(FValid(DSqueeze2(F.LRO, F2))), PC(S(Last(DSqueeze2(F.LRO, F2))))). + + lemma equiv_ideal (D <: DISTINGUISHER{-SLCommon.C, -C, -IF, -S, -F.FRO, -F2.RO, -F2.FRO, -BIRO.IRO, -BIRO2.IRO}) &m: + Pr[SLCommon.IdealIndif(IF,S,SLCommon.DRestr(A(D))).main() @ &m : res] = + Pr[IdealIndif(BIRO.IRO,SimLast(S),DRestr(D)).main() @ &m : res]. + proof. + have->:Pr[SLCommon.IdealIndif(IF, S, SLCommon.DRestr(A(D))).main() @ &m : res] + = Pr[SLCommon.IdealIndif(IF, S, A(D)).main() @ &m : res]. + + by byequiv(ideal_equiv D)=>//=. + have<-:Pr[L2(D,F.RO).distinguish() @ &m : res] = + Pr[SLCommon.IdealIndif(IF,S,A(D)).main() @ &m : res]. + + by byequiv(ideal_equiv2 D). + have->:Pr[L2(D, F.RO).distinguish() @ &m : res] = + Pr[L2(D,F.LRO).distinguish() @ &m : res]. + + byequiv=>//=;proc;sp;inline*;sp;wp. + call(F.FullEager.RO_LRO_D (D2(D)) _);auto. + by move=> _; exact/dunifin_ll. + have->:Pr[IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res] = + Pr[L4(D,F.LRO,F2.LRO).distinguish() @ &m : res]. + + by rewrite eq_sym;byequiv(equiv_L4_ideal D)=>//=. + have<-:Pr[L4(D, F.RO, F2.RO).distinguish() @ &m : res] = + Pr[L4(D,F.LRO,F2.LRO).distinguish() @ &m : res]. + + have->:Pr[L4(D, F.RO, F2.RO).distinguish() @ &m : res] = + Pr[L4(D,F.LRO, F2.RO).distinguish() @ &m : res]. + - byequiv=>//=;proc;sp;inline*;sp;wp. + call(F.FullEager.RO_LRO_D (D5(D)) _); auto. + by move=> _; exact/dunifin_ll. + byequiv=>//=;proc;sp;inline*;sp;wp. + call(F2.FullEager.RO_LRO_D (D6(D)) _); auto. + by move=> _; exact/dunifin_ll. + have<-:Pr[L3(D, F.RO).distinguish() @ &m : res] = + Pr[L4(D, F.RO, F2.RO).distinguish() @ &m : res]. + + by byequiv(equiv_L3_L4 D)=>//=. + have<-:Pr[L(D, F.RO).distinguish() @ &m : res] = + Pr[L3(D, F.RO).distinguish() @ &m : res]. + + by byequiv(Ideal_equiv3 D). + have->:Pr[L(D, F.RO).distinguish() @ &m : res] = + Pr[L(D,F.LRO).distinguish() @ &m : res]. + + byequiv=>//=;proc;sp;inline*;sp;wp. + call(F.FullEager.RO_LRO_D (D3(D)) _); auto. + by move=> _; exact/dunifin_ll. + rewrite eq_sym. + by byequiv(Ideal_equiv_valid D). + qed. + +end section Ideal. + + + (* Real part *) + + +section Real. + + inductive m_p (m : (state, state) fmap) (p : (block list, state) fmap) = + | IND_M_P of (p.[[]] = Some (b0, c0)) + & (forall l, l \in p => forall i, 0 <= i < size l => + exists b c, p.[take i l] = Some (b,c) /\ + m.[(b +^ nth witness l i, c)] = p.[take (i+1) l]). + + inductive INV_Real + (c1 c2 : int) + (m mi : (state, state) fmap) + (p : (block list, state) fmap) = + | INV_real of (c1 <= c2) + & (m_p m p) + & (invm m mi). + + local lemma INV_Real_incr c1 c2 m mi p : + INV_Real c1 c2 m mi p => + INV_Real (c1 + 1) (c2 + 1) m mi p. + proof. by case;progress;split=>//=/#. qed. + + local lemma INV_Real_addm_mi c1 c2 m mi p x y : + INV_Real c1 c2 m mi p => + ! x \in m => + ! rng m y => + INV_Real c1 c2 m.[x <- y] mi.[y <- x] p. + proof. + case=> H_c1c2 H_m_p H_invm H_x_dom H_y_rng;split=>//=. + + split;case:H_m_p=>//=; + smt(get_setE domE oget_some take_oversize size_take take_take). + exact/invm_set. + qed. + + local lemma invmC' (m mi : (state, state) fmap) : + invm m mi => invm mi m. + proof. by rewrite /#. qed. + + local lemma invmC (m mi : (state, state) fmap) : + invm m mi <=> invm mi m. + proof. by split;exact invmC'. qed. + + local lemma invm_dom_rng (m mi : (state, state) fmap) : + invm m mi => dom m = rng mi. + proof. + move=>h; rewrite fun_ext=> x; rewrite domE rngE /= eq_iff; have h2 := h x; split. + + move=> m_x_not_None; exists (oget m.[x]); rewrite -h2; move: m_x_not_None. + by case: (m.[x]). + by move=> [] a; rewrite -h2 => ->. + qed. + + local lemma all_prefixes_of_INV_real c1 c2 m mi p: + INV_Real c1 c2 m mi p => + all_prefixes p. + proof. + move=>[]_[]Hp0 Hmp1 _ l H_dom i. + smt(take_le0 take_oversize size_take take_take take_size nth_take domE). + qed. + + local lemma lemma2' c1 c2 m mi p bl i sa sc: + INV_Real c1 c2 m mi p => + 1 < i => + valid bl => + (sa,sc) \in m => + ! (format bl i) \in p => + p.[format bl (i-1)] = Some (sa,sc) => + INV_Real c1 c2 m mi p.[format bl i <- oget m.[(sa,sc)]]. + proof. + move=>inv0 h1i h_valid H_dom_m H_dom_p H_p_val. + split;have[]//=_[] hmp0 hmp1 hinvm:=inv0;split=>//=. + + by rewrite get_setE;smt(size_cat size_nseq size_ge0). + + move=>l;rewrite mem_set;case;1:smt(all_prefixes_of_INV_real get_setE). + move=>->>j[]hj0 hjsize;rewrite get_setE/=. + have:=hmp1 (format bl (i - 1));rewrite domE H_p_val/==>help. + have:=hjsize;rewrite !size_cat !size_nseq/=!ler_maxr 1:/#=>hjsizei. + have->/=:!take j (format bl i) = format bl i by smt(size_take). + have h:forall k, 0 <= k <= size bl + i - 2 => + take k (format bl (i - 1)) = take k (format bl i). + * move=>k[]hk0 hkjS;rewrite !take_cat;case(k//=hksize;congr. + apply (eq_from_nth witness);1:rewrite!size_take//=1,2:/#!size_nseq!ler_maxr/#. + rewrite!size_take//=1:/#!size_nseq!ler_maxr 1:/#. + pose o:=if _ then _ else _;have->/={o}:o = k - size bl by smt(). + by progress;rewrite!nth_take//= 1,2:/# !nth_nseq//=/#. + case(j < size bl + i - 2)=>hj. + - have:=help j _;1:smt(size_cat size_nseq). + move=>[]b c[]. + have->:nth witness (format bl (i - 1)) j = nth witness (format bl i) j. + + by rewrite-(nth_take witness (j+1)) 1,2:/# eq_sym -(nth_take witness (j+1)) 1,2:/# !h//=/#. + rewrite h 1:/# h 1:/# => -> h';exists b c=>//=;rewrite h'/=get_setE/=. + smt(size_take size_cat size_nseq). + have->>/=:j = size (format bl (i-1)) by smt(size_cat size_nseq). + rewrite get_setE/=. + have h':size (format bl (i-1)) = size bl + i - 2 by smt(size_cat size_nseq). + rewrite h'/=. + have h'':(size bl + i - 1) = size (format bl i) by smt(size_cat size_nseq). + rewrite h'' take_size/=-h 1:/# -h' take_size. + rewrite nth_cat h';have->/=:! size bl + i - 2 < size bl by smt(). + by rewrite nth_nseq 1:/#; exists sa sc; smt(Block.WRing.AddMonoid.addm0 domE). + qed. + + local lemma take_nseq (a : 'a) i j : + take j (nseq i a) = if j <= i then nseq j a else nseq i a. + proof. + case(0 <= j)=>hj0;last first. + + rewrite take_le0 1:/#;smt(nseq0_le). + case(j <= i)=>hij//=;last smt(take_oversize size_nseq). + apply(eq_from_nth witness). + + smt(size_take size_nseq). + smt(size_nseq size_take nth_take nth_nseq). + qed. + + local lemma take_format (bl : block list) n i : + 0 < n => + 0 <= i < size bl + n => + take i (format bl n) = + if i <= size bl then take i bl else format bl (i - size bl + 1). + proof. + move=>Hn0[]Hi0 Hisize;rewrite take_cat take_nseq. + case(i < size bl)=>//=[/#|H_isize']. + have->/=:i - size bl <= n - 1 by smt(). + case(i = size bl)=>[->>|H_isize'']//=;1:by rewrite nseq0 take_size cats0. + smt(). + qed. + + + local lemma equiv_sponge (D <: DISTINGUISHER {-P, -Redo, -C, -SLCommon.C}) : + equiv [ GReal(A(D)).main + ~ NIndif(Squeeze(SqueezelessSponge(P)),P,DRestr(D)).main + : ={glob D} ==> ={res, glob D, glob P, C.c} /\ SLCommon.C.c{1} <= C.c{2} <= max_size]. + proof. + proc;inline*;sp;wp. + call(: ={Redo.prefixes, glob P, C.c} /\ C.c{1} <= max_size /\ + INV_Real SLCommon.C.c{1} C.c{2} Perm.m{1} Perm.mi{1} Redo.prefixes{1});auto;last first. + + progress. + + exact max_ge0. + + by split=>//=;1:split;smt(mem_empty in_fset0 mem_set get_setE). + by case:H2=>//=. + + by proc;inline*;auto;sp;if;auto;sp;if;auto; + smt(INV_Real_addm_mi INV_Real_incr supp_dexcepted). + + proc;inline*;auto;sp;if;auto;sp;if;auto;progress. + + apply INV_Real_incr=>//=. + apply INV_Real_addm_mi=>//=. + + case:H0=>H_c H_m_p H_invm;rewrite (invm_dom_rng _ _ H_invm)//=. + by move:H3;rewrite supp_dexcepted. + case:H0=>H_c H_m_p H_invm;have<-//:=(invm_dom_rng Perm.mi{2} Perm.m{2}). + by rewrite invmC. + + exact INV_Real_incr. + + proc;inline*;sp;if;auto. + swap 6;wp;sp=>/=;if;auto;last by progress;split;case:H0=>//=;smt(size_ge0). + conseq(: p{2} = bl{2} /\ n{2} = nb{2} /\ lres{2} = [] /\ b{2} = b0 /\ + i{2} = 0 /\ p{1} = bl{1} /\ n{1} = nb{1} /\ lres{1} = [] /\ b{1} = b0 /\ + i{1} = 0 /\ z{2} = [] /\ z{1} = [] /\ ={bl, nb} /\ ={Redo.prefixes} /\ + ={Perm.mi, Perm.m} /\ ={C.c} /\ + INV_Real SLCommon.C.c{1} C.c{2} Perm.m{1} Perm.mi{1} Redo.prefixes{1} /\ + C.c{1} + size bl{1} + max (nb{1} - 1) 0 <= max_size /\ valid p{1} + ==> ={lres} /\ ={Redo.prefixes} /\ ={Perm.mi, Perm.m} /\ + C.c{1} + size bl{1} + max (nb{1} - 1) 0 = + C.c{2} + size bl{2} + max (nb{2} - 1) 0 /\ + INV_Real SLCommon.C.c{1} (C.c{2} + size bl{2} + max (nb{2} - 1) 0) + Perm.m{1} Perm.mi{1} Redo.prefixes{1});progress. + sp. + seq 2 2:(={i,n,p,lres,nb,bl,b,glob P,glob C,glob Redo} + /\ INV_Real SLCommon.C.c{1} (C.c{2} + size bl{2}) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ (n,p){1} = (nb,bl){1} /\ lres{1} = [] /\ i{1} = 0 + /\ valid p{1} + /\ Redo.prefixes.[p]{1} = Some (b,sc){1}). + + exists* Redo.prefixes{1},SLCommon.C.c{1};elim* => pref count/=. + wp;conseq(:_==> ={i0,p0,i,p,n,nb,bl,sa,lres,C.c,glob Redo,glob Perm} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p0{1} = p{1} /\ i0{1} = size p{1} + /\ Redo.prefixes{1}.[take i0{1} p{1}] = Some (sa{1},sc{1}) + /\ INV_Real count C.c{1} Perm.m{1} Perm.mi{1} pref + /\ (forall l, l \in Redo.prefixes{1} => + l \in pref \/ (exists j, 0 < j <= i0{2} /\ l = take j p{1})) + /\ (forall l, l \in pref => pref.[l] = Redo.prefixes{1}.[l]) + /\ SLCommon.C.c{1} <= count + i0{1} <= C.c{1} + i0{1} + /\ (forall j, 0 <= j < i0{1} => + exists b c, Redo.prefixes{1}.[take j p{1}] = Some (b,c) /\ + Perm.m{1}.[(b +^ nth witness p{1} j, c)] = + Redo.prefixes{1}.[take (j+1) p{1}])); + progress. + - have inv0:=H3;have[]h_c1c2[]Hmp0 Hmp1 Hinvm:=inv0;split=>//=. + - case:inv0;smt(size_ge0). + split=>//=. + - smt(domE). + - move=>l H_dom_R i []Hi0 Hisize;have:=H4 l H_dom_R. + case(l \in Redo.prefixes{2})=>H_in_pref//=. + * have:=Hmp1 l H_in_pref i _;rewrite//=. + rewrite ?H5//=;1:smt(domE). + case(i+1 < size l)=>h;1:smt(domE). + by rewrite take_oversize 1:/#. + move=>[]j[][]hj0 hjsize ->>. + have:=Hisize;rewrite size_take 1:/#. + pose k:=if _ then _ else _. + have: k = j by smt(). + move: k=> /> Hij. + by rewrite!take_take!minrE 1:nth_take 1,2:/#;smt(domE). + - smt(get_setE oget_some domE take_oversize). + while( ={i0,p0,i,p,n,nb,bl,sa,sc,lres,C.c,glob Redo,glob Perm} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p0{1} = p{1} /\ 0 <= i0{1} <= size p{1} + /\ Redo.prefixes{1}.[take i0{1} p{1}] = Some (sa{1},sc{1}) + /\ INV_Real count C.c{1} Perm.m{1} Perm.mi{1} pref + /\ (forall l, l \in Redo.prefixes{1} => + l \in pref \/ (exists j, 0 < j <= i0{2} /\ l = take j p{1})) + /\ (forall l, l \in pref => pref.[l] = Redo.prefixes{1}.[l]) + /\ SLCommon.C.c{1} <= count + i0{1} <= C.c{1} + i0{1} + /\ (i0{1} < size p0{1} => + take (i0{1}+1) p{1} \in Redo.prefixes{1} => + Redo.prefixes{1} = pref) + /\ all_prefixes Redo.prefixes{1} + /\ (forall j, 0 <= j < i0{1} => + exists b c, Redo.prefixes{1}.[take j p{1}] = Some (b,c) /\ + Perm.m{1}.[(b +^ nth witness p{1} j, c)] = + Redo.prefixes{1}.[take (j+1) p{1}]));last first. + + auto;progress. + - exact size_ge0. + - by rewrite take0;have[]_[]->//=:=H. + - smt(). + - by have[]->//=:=H. + - smt(all_prefixes_of_INV_real). + - smt(). + - smt(). + if;auto;progress. + - smt(). + - smt(). + - smt(domE). + - smt(domE). + - smt(). + - smt(). + - smt(all_prefixes_of_INV_real domE take_take size_take). + - case(j < i0{2})=>hj;1:smt(). + have<<-/=:j = i0{2} by smt(). + have->>:=H7 H10 H12. + have[]_[]hmp0 hmp1 _:=H2. + have[]b3 c3:=hmp1 _ H12 j _;1:smt(size_take). + smt(take_take nth_take size_take). + sp;if;auto;progress. + - smt(). + - smt(). + - smt(get_setE domE). + - rewrite INV_Real_addm_mi//=;smt(supp_dexcepted). + - smt(mem_set). + - smt(get_setE domE). + - smt(). + - smt(). + - move:H17;apply absurd=>//=_;rewrite mem_set. + pose x:=_ = _;have->/={x}:x=false by smt(size_take). + move:H12;apply absurd=>//= hpref. + have:= H8 _ hpref (i0{2}+1). + smt(mem_set take_take size_take). + - move=>l;rewrite!mem_set;case=>[H_dom i|->>]/=. + * by rewrite mem_set;smt(). + move=>j; case(0 <= j)=>hj0; rewrite mem_set. + * case: (j <= i0{2}) => hjmax; 2:smt(take_oversize size_take). + left; have-> : take j (take (i0{2}+1) bl{2}) = take j (take i0{2} bl{2}). + * by rewrite 2!take_take !minrE /#. + by apply H8; rewrite domE H1. + rewrite take_le0 1:/#; left. + by rewrite-(take0 (take i0{2} bl{2})) H8 domE H1. + - smt(get_setE domE mem_set). + - smt(get_setE domE). + - smt(). + - smt(get_setE domE). + - smt(mem_set). + - smt(get_setE domE). + - smt(). + - smt(). + - move:H15;apply absurd=>//=_;rewrite mem_set. + pose x:=_ = _;have->/={x}:x=false by smt(size_take). + move:H12;apply absurd=>//=. + have:=take_take bl{2}(i0{2} + 1)(i0{2} + 1 + 1). + rewrite minrE (: i0{2} + 1 <= i0{2} + 1 + 1) 1:/#=><-h. + by rewrite (H8 _ h). + - move=>l;rewrite!mem_set;case=>[H_dom|->>]/=;1:smt(mem_set). + move=>j;rewrite mem_set. + case(0 <= j)=>hj0; last first. + * rewrite take_le0 1:/#; left. + by rewrite-(take0 (take i0{2} bl{2})) H8 domE H1. + case(j < i0{2} + 1)=>hjiS;2:smt(domE take_take). + rewrite take_take/min hjiS//=;left. + have:=(take_take bl{2} j i0{2}). + rewrite minrE (: j <= i0{2}) 1:/#=><-. + smt(all_prefixes_of_INV_real domE). + - smt(get_setE domE mem_set). + sp;case(0 < n{1});last first. + - rcondf{1}1;2:rcondf{2}1;auto;1:smt(). + splitwhile{1} 1 : i + 1 < n;splitwhile{2} 1 : i + 1 < n. + rcondt{1}2;2:rcondt{2}2;auto;progress. + + while(i < n);auto. + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + + while(i < n);auto. + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + rcondf{1}4;2:rcondf{2}4;auto. + + while(i < n);auto;2:smt(). + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + + while(i < n);auto;2:smt(). + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + rcondf{1}4;2:rcondf{2}4;1,2:auto. + + while(i < n);auto;2:smt(). + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + + while(i < n);auto;2:smt(). + by sp;if;auto;sp;while(i < n);auto;if;auto;sp;if;auto. + conseq(:_==> ={i,n,p,lres,nb,bl,b,glob P,glob C,glob Redo} + /\ INV_Real SLCommon.C.c{1} (C.c{2} + size bl{2} + i{1} - 1) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ i{1} = n{1});1:smt();wp. + conseq(:_==> ={i,n,p,lres,nb,bl,b,glob P,glob C,glob Redo} + /\ INV_Real SLCommon.C.c{1} (C.c{2} + size bl{2} + i{1}) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ i{1}+1 = n{1});1:smt(). + while(={i,n,p,lres,nb,bl,b,glob P,glob C,glob Redo} + /\ INV_Real SLCommon.C.c{1} (C.c{2} + size bl{2} + i{1}) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ (n,p){1} = (nb,bl){1} /\ 0 < i{1}+1 <= n{1} + /\ valid p{1} + /\ (exists c2, Redo.prefixes.[format p (i+1)]{1} = Some (b,c2){1})); + last by auto;smt(nseq0 cats0). + sp;rcondt{1}1;2:rcondt{2}1;auto. + sp. + splitwhile{1} 1 : i1 < size p1 - 1;splitwhile{2} 1 : i1 < size p1 - 1. + rcondt{1}2;2:rcondt{2}2;1,2:by auto; + while(i1 < size p1);auto;1:if;2:(sp;if);auto;smt(size_cat size_nseq size_ge0). + rcondf{1}4;2:rcondf{2}4;1,2:by auto; + seq 1 : (i1 = size p1 - 1);1:(auto; + while(i1 < size p1);auto;1:if;2:(sp;if);auto;smt(size_cat size_nseq size_ge0)); + if;sp;2:if;auto;smt(size_cat size_nseq size_ge0). + wp=>//=. + wp;conseq(:_==> ={sa0,sc0,glob Redo,glob Perm} + /\ INV_Real SLCommon.C.c{1} (C.c{1} + size bl{2} + i{1}) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ (format p{1} i{1} \in Redo.prefixes{1}) + /\ exists (c2 : capacity), Redo.prefixes{1}.[format p{1} (i{1}+1)] = Some (sa0{1}, c2));progress. + + smt(size_ge0). + + smt(size_ge0). + + smt(). + seq 1 1 : (={nb,bl,n,p,p1,i,i1,lres,sa0,sc0,C.c,glob Redo,glob Perm} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p1{1} = format p{1} (i{1}+1) + /\ 1 <= i{1} < n{1} /\ valid p{1} /\ i1{1} = size p1{1} - 1 + /\ Redo.prefixes{1}.[format p{1} i{1}] = Some (sa0{1},sc0{1}) + /\ INV_Real SLCommon.C.c{1} (C.c{1} + size bl{2} + i{1} - 1) Perm.m{1} Perm.mi{1} + Redo.prefixes{1});last first. + + if;auto;progress. + - by split;case:H3=>//=;smt(). + - by rewrite domE H2//=. + - move:H4;rewrite take_size /= domE=> h. + exists (oget Redo.prefixes{2}.[format bl{2} (i{2} + 1)]).`2; move: h. + by case: (Redo.prefixes{2}.[format bl{2} (i{2} + 1)]); smt(). + sp;if;auto;progress. + - move:H4 H5;rewrite!get_setE/= nth_last/=take_size. + rewrite last_cat last_nseq 1:/# Block.WRing.addr0;progress. + have//=:=lemma2'(SLCommon.C.c{1} + 1)(C.c{2} + size bl{2} + i{2}) + Perm.m{2}.[(sa0_R, sc0{2}) <- y2L] Perm.mi{2}.[y2L <- (sa0_R, sc0{2})] + Redo.prefixes{2} bl{2} (i{2}+1) sa0_R sc0{2}. + rewrite H1/=!mem_set/=H4/=H2/=get_setE/=. + have->->//=:y2L = (y2L.`1, y2L.`2);1,-1:smt(). + rewrite INV_Real_addm_mi//=;2:smt(supp_dexcepted). + by have:=H3=>hinv0;split;case:hinv0=>//=/#. + - by rewrite mem_set//=take_size domE H2. + - by rewrite!get_setE take_size/=;smt(). + - move:H4 H5;rewrite nth_last take_size. + rewrite last_cat last_nseq 1:/# Block.WRing.addr0;progress. + pose a:=(_, _);have->/={a}:a = oget Perm.m{2}.[(sa0_R, sc0{2})] by smt(). + apply lemma2'=>//=;first have:=H3=>hinv0;split;case:hinv0=>//=/#. + smt(). + - by rewrite mem_set//=take_size;smt(domE). + - by rewrite!get_setE/=take_size/=;smt(). + alias{1} 1 pref = Redo.prefixes;sp;alias{1} 1 count = SLCommon.C.c. + alias{1} 1 pm = Perm.m;sp;alias{1} 1 pmi = Perm.mi;sp. + conseq(:_==> ={nb,bl,n,p,p1,i,i1,lres,sa0,sc0,C.c,glob Redo,glob Perm} + /\ pmi{1} = Perm.mi{1} /\ pm{1} = Perm.m{1} + /\ pref{1} = Redo.prefixes{1} /\ SLCommon.C.c{1} = count{1} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p1{1} = format p{1} (i{1}+1) + /\ i1{1} = size p1{1} - 1 + /\ Redo.prefixes{1}.[format p{1} (i1{1} - size p{1} + 1)] = + Some (sa0{1}, sc0{1}));progress. + + smt(). + + by move: H8; rewrite size_cat size_nseq /= ler_maxr /#. + + move:H8;rewrite size_cat size_nseq/=/max H0/=;smt(). + splitwhile{1}1:i1 < size p;splitwhile{2}1:i1 < size p. + while(={nb,bl,n,p,p1,i,i1,lres,sa0,sc0,C.c,glob Redo,glob Perm} + /\ INV_Real SLCommon.C.c{1} (C.c{1} + size bl{2} + i{1} - 1) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ pmi{1} = Perm.mi{1} /\ pm{1} = Perm.m{1} + /\ pref{1} = Redo.prefixes{1} /\ SLCommon.C.c{1} = count{1} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p1{1} = format p{1} (i{1}+1) + /\ (format p{1} i{1} \in Redo.prefixes{1}) + /\ size p{1} <= i1{1} <= size p1{1} - 1 /\ valid p{1} + /\ Redo.prefixes{1}.[format p{1} (i1{1} - size p{1} + 1)] = + Some (sa0{1}, sc0{1})). + + rcondt{1}1;2:rcondt{2}1;auto;progress. + + have->:take (i1{m} + 1) (format bl{m} (i{m} + 1)) = + take (i1{m} + 1) (format bl{m} i{m});2:smt(all_prefixes_of_INV_real). + smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + + have->:take (i1{hr} + 1) (format bl{hr} (i{hr} + 1)) = + take (i1{hr} + 1) (format bl{hr} i{hr});2:smt(all_prefixes_of_INV_real). + smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + + smt(). + + smt(size_cat size_nseq). + + have->:take (i1{2} + 1) (format bl{2} (i{2} + 1)) = + take (i1{2} + 1) (format bl{2} i{2}). + - smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + have->:format bl{2} (i1{2} + 1 - size bl{2} + 1) = + take (i1{2} + 1) (format bl{2} i{2}). + - smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + have all_pref:=all_prefixes_of_INV_real _ _ _ _ _ H. + by have:=all_pref _ H0 (i1{2}+1); rewrite domE; smt(). + conseq(:_==> ={nb,bl,n,p,p1,i,i1,lres,sa0,sc0,C.c,glob Redo,glob Perm} + /\ INV_Real SLCommon.C.c{1} (C.c{1} + size bl{2} + i{1} - 1) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ pmi{1} = Perm.mi{1} /\ pm{1} = Perm.m{1} + /\ pref{1} = Redo.prefixes{1} /\ SLCommon.C.c{1} = count{1} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p1{1} = format p{1} (i{1}+1) + /\ (format p{1} i{1} \in Redo.prefixes{1}) + /\ i1{1} = size p{1} /\ valid p{1} + /\ Redo.prefixes{1}.[take i1{1} p{1}] = Some (sa0{1}, sc0{1})); + 1:smt(size_cat size_nseq nseq0 cats0 take_size). + while(={nb,bl,n,p,p1,i,i1,lres,sa0,sc0,C.c,glob Redo,glob Perm} + /\ INV_Real SLCommon.C.c{1} (C.c{1} + size bl{2} + i{1} - 1) + Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ pmi{1} = Perm.mi{1} /\ pm{1} = Perm.m{1} + /\ pref{1} = Redo.prefixes{1} /\ SLCommon.C.c{1} = count{1} + /\ n{1} = nb{1} /\ p{1} = bl{1} /\ p1{1} = format p{1} (i{1}+1) + /\ (format p{1} i{1} \in Redo.prefixes{1}) + /\ 0 <= i1{1} <= size p{1} /\ valid p{1} + /\ Redo.prefixes{1}.[take i1{1} p{1}] = Some (sa0{1}, sc0{1}));last first. + + auto;progress. + - smt(). + - have[]_[]:=H;smt(domE). + - exact size_ge0. + - have[]_[]:=H;smt(domE take0). + - smt(size_cat size_nseq). + rcondt{1}1;2:rcondt{2}1;auto;progress. + - have->:take (i1{m} + 1) (format bl{m} (i{m} + 1)) = + take (i1{m} + 1) (format bl{m} i{m});2:smt(all_prefixes_of_INV_real). + smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + - have->:take (i1{hr} + 1) (format bl{hr} (i{hr} + 1)) = + take (i1{hr} + 1) (format bl{hr} i{hr});2:smt(all_prefixes_of_INV_real). + smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + - smt(). + - smt(). + - have->:take (i1{2} + 1) (format bl{2} (i{2} + 1)) = + take (i1{2} + 1) (format bl{2} i{2}) + by smt(take_format size_ge0 size_eq0 valid_spec size_cat size_nseq). + have->:take (i1{2} + 1) bl{2} = + take (i1{2} + 1) (format bl{2} i{2}) + by smt(take_cat take_le0 cats0). + smt(all_prefixes_of_INV_real). + qed. + + + local lemma lemma4 c c' m mi p bl i sa sc: + INV_Real c c' m mi p => + 0 < i => + p.[format bl i] = Some (sa,sc) => + format bl (i+1) \in p => + p.[format bl (i+1)] = m.[(sa,sc)]. + proof. + move=>inv0 H_i0 H_p_i H_dom_iS. + have[]_[]_ hmp1 _ :=inv0. + have:=hmp1 (format bl (i+1)) H_dom_iS=>help. + have:=help (size (format bl i)) _;1:smt(size_ge0 size_cat size_nseq). + move=>[]b3 c3;rewrite!take_format;..4:smt(size_ge0 size_cat size_nseq). + have->/=:!size (format bl i) + 1 <= size bl by smt(size_cat size_nseq size_ge0). + rewrite nth_cat. + have->/=:!size (format bl i) < size bl by smt(size_cat size_ge0). + rewrite nth_nseq 1:size_cat 1:size_nseq 1:/#. + pose x:=if _ then _ else _;have->/={x}:x = format bl i. + + rewrite/x;case(i = 1)=>//=[->>|hi1]. + - by rewrite/format/=nseq0 cats0//=take_size. + by rewrite size_cat size_nseq/#. + pose x:=List.size _ + 1 - List.size _ + 1;have->/={x}:x=i+1 + by rewrite/x size_cat size_nseq;smt(). + rewrite H_p_i=>[]/=[][]->>->>. + by rewrite Block.WRing.addr0=>H_pm;rewrite H_pm/=. + qed. + + local lemma lemma_3 c1 c2 m mi p bl b (sa:block) sc: + INV_Real c1 c2 m mi p => + (sa +^ b,sc) \in m => + ! rcons bl b \in p => + p.[bl] = Some (sa,sc) => + INV_Real c1 c2 m mi p.[rcons bl b <- oget m.[(sa +^ b,sc)]]. + proof. + move=>inv0 H_dom_m H_dom_p H_p_val. + split;have[]//=_[] hmp0 hmp1 hinvm:=inv0;split=>//=. + + by rewrite get_setE;smt(size_cat size_nseq size_ge0). + + move=>l;rewrite mem_set;case;1:smt(all_prefixes_of_INV_real get_setE). + move=>->>j[]hj0 hjsize;rewrite get_setE/=. + have:=hmp1 bl;rewrite domE H_p_val/==>help. + have->/=:!take j (rcons bl b) = rcons bl b by smt(size_take). + move:hjsize;rewrite size_rcons=>hjsize. + rewrite-cats1 !take_cat. + pose x := if _ then _ else _;have->/={x}: x = take j bl by smt(take_le0 cats0 take_size). + rewrite nth_cat. + case(j < size bl)=>//=hj;last first. + + have->>/=:j = size bl by smt(). + by rewrite take_size H_p_val/=;exists sa sc=>//=;smt(get_setE). + have->/=:j + 1 - size bl <= 0 by smt(). + rewrite cats0. + pose x := if _ then _ else _;have->/={x}: x = take (j+1) bl by smt(take_le0 cats0 take_size). + have:=hmp1 bl;rewrite domE H_p_val/==>hep. + have:=hep j _;rewrite//=;smt(get_setE size_cat size_take). + qed. + + + + local lemma squeeze_squeezeless (D <: DISTINGUISHER {-P, -Redo, -C, -SLCommon.C}) : + equiv [ NIndif(Squeeze(SqueezelessSponge(P)),P,DRestr(D)).main + ~ RealIndif(Sponge,P,DRestr(D)).main + : ={glob D} ==> ={res, glob P, glob D, C.c} /\ C.c{1} <= max_size]. + proof. + proc;inline*;sp;wp. + call(: ={glob Perm,C.c} /\ C.c{1} <= max_size + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1});auto;last first. + + progress. + + exact max_ge0. + split=>//=;1:split=>//=;smt(get_setE mem_empty emptyE in_fset0 mem_set). + + proc;inline*;auto;sp;if;auto;sp;if;auto;progress. + - by rewrite INV_Real_addm_mi;2..:smt(supp_dexcepted);split;case:H0=>//=;smt(). + - by split;case:H0=>//=;smt(). + + proc;inline*;auto;sp;if;auto;sp;if;auto;progress. + - rewrite INV_Real_addm_mi;1: by split;case:H0=>//=;smt(). + * case:H0;smt(invm_dom_rng invmC supp_dexcepted). + case:H0;smt(invm_dom_rng invmC supp_dexcepted). + - by split;case:H0=>//=;smt(). + proc;inline*;sp;auto;if;auto;sp;if;auto; + last by progress;split;case:H0=>//=;smt(size_ge0). + conseq(: (exists (c_R : int), + C.c{2} = c_R + size bl{2} + max (nb{2} - 1) 0 /\ xs{2} = bl{2} /\ + n{2} = nb{2} /\ z0{2} = [] /\ sc{2} = c0 /\ sa{2} = b0 /\ i{2} = 0 /\ + exists (c_L : int), C.c{1} = c_L + size bl{1} + max (nb{1} - 1) 0 /\ + p{1} = bl{1} /\ n{1} = nb{1} /\ lres{1} = [] /\ b{1} = b0 /\ + i{1} = 0 /\ z{2} = [] /\ z{1} = [] /\ ={bl, nb} /\ + ={Perm.mi, Perm.m} /\ c_L = c_R /\ + INV_Real 0 c_L Perm.m{1} Perm.mi{1} Redo.prefixes{1} /\ valid p{1}) + ==> lres{1} = z0{2} /\ ={Perm.mi, Perm.m} /\ ={C.c} /\ + INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1});1,2:smt(). + sp. + seq 2 1 : (={glob P, i, n, C.c,sa,sc} + /\ b{1} = sa{2} /\ Redo.prefixes.[p]{1} = Some (sa,sc){2} + /\ lres{1} = z0{2} /\ i{1} = 0 /\ valid p{1} + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1}). + + conseq(:_==> ={glob P, n, C.c,sa,sc} /\ b{1} = sa{2} /\ i0{1} = size p0{1} + /\ Redo.prefixes{1}.[take i0{1} p0{1}] = Some (sa{1}, sc{1}) + /\ lres{1} = z0{2} /\ xs{2} = drop i0{1} p0{1} + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1});1:smt(take_size drop_size). + wp;while(={glob P, n, C.c,sa,sc} /\ sa{1} = sa{2} /\ sc{1} = sc{2} + /\ 0 <= i0{1} <= size p0{1} + /\ Redo.prefixes{1}.[take i0{1} p0{1}] = Some (sa{1}, sc{1}) + /\ lres{1} = z0{2} /\ xs{2} = drop i0{1} p0{1} + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1}). + + if{1};auto. + + sp;rcondf{2}1;auto;progress. + + rewrite head_nth nth_drop//=. + have[]_[]_ hmp1 _ :=H2;have:=hmp1 _ H5 i0{m} _;1:smt(size_take). + move=>[]b3 c3;rewrite!take_take!nth_take 1,2:/# !minrE //= (: i0{m} <= i0{m} + 1) 1:/#. + rewrite H1=>//=[][][]->>->>. + by rewrite nth_onth (onth_nth b0)//=;smt(domE). + + rewrite head_nth nth_drop//=. + have[]_[]_ hmp1 _ :=H2;have:=hmp1 _ H5 i0{1} _;1:smt(size_take). + move=>[]b3 c3;rewrite!take_take!nth_take 1,2:/# !minrE//= (: i0{1} <= i0{1} + 1) 1:/#. + rewrite H1=>//=[][][]->>->>. + by rewrite nth_onth (onth_nth b0)//=;smt(domE). + + rewrite head_nth nth_drop//=. + have[]_[]_ hmp1 _ :=H2;have:=hmp1 _ H5 i0{1} _;1:smt(size_take). + move=>[]b3 c3;rewrite!take_take!nth_take 1,2:/# !minrE //= (: i0{1} <= i0{1} + 1) 1:/#. + rewrite H1=>//=[][][]->>->>. + by rewrite nth_onth (onth_nth b0)//=;smt(domE). + + rewrite head_nth nth_drop//=. + have[]_[]_ hmp1 _ :=H2;have:=hmp1 _ H5 i0{1} _;1:smt(size_take). + move=>[]b3 c3;rewrite!take_take!nth_take 1,2:/# !minrE //= (: i0{1} <= i0{1} + 1) 1:/#. + rewrite H1=>//=[][][]->>->>. + by rewrite nth_onth (onth_nth b0)//=;smt(domE). + + rewrite head_nth nth_drop//=. + have[]_[]_ hmp1 _ :=H2;have:=hmp1 _ H5 i0{1} _;1:smt(size_take). + move=>[]b3 c3;rewrite!take_take!nth_take 1,2:/# !minrE //= (: i0{1} <= i0{1} + 1) 1:/#. + rewrite H1=>//=[][][]->>->>. + by rewrite nth_onth (onth_nth b0)//=;smt(domE). + + smt(). + + smt(). + + smt(). + + smt(behead_drop drop_add). + + smt(size_drop size_eq0). + + smt(size_drop size_eq0). + sp=>//=. + if;auto;progress. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth witness)//=. + + by move:H6;rewrite head_nth nth_drop //=nth_onth (onth_nth witness)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + smt(). + + smt(). + + by rewrite get_setE/=. + + by rewrite behead_drop drop_add. + + rewrite!get_setE/=. + have:=lemma_3 0 C.c{2}Perm.m{2}.[(sa{2} +^ nth witness p0{1} i0{1}, sc{2}) <- yL] + Perm.mi{2}.[yL <- (sa{2} +^ nth witness p0{1} i0{1}, sc{2})] Redo.prefixes{1} + (take i0{1} p0{1}) (nth witness p0{1} i0{1}) sa{2} sc{2}. + rewrite!mem_set/=-take_nth//=H5/=H1/=get_setE/=. + have->->//=:(yL.`1, yL.`2) = yL by smt(). + rewrite INV_Real_addm_mi=>//=;smt(supp_dexcepted). + + smt(size_drop size_eq0). + + smt(size_drop size_eq0). + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + by rewrite head_nth nth_drop //=nth_onth (onth_nth b0)//=. + + smt(). + + smt(). + + by rewrite get_setE. + + by rewrite behead_drop drop_add. + + rewrite(take_nth witness)//=. + have:=lemma_3 0 C.c{2} Perm.m{2} Perm.mi{2} Redo.prefixes{1} + (take i0{1} p0{1}) (nth witness p0{1} i0{1}) sa{2} sc{2}. + by rewrite-take_nth//= H5/=H1/=H2/=H6/=;smt(). + + smt(size_drop size_eq0). + + smt(size_drop size_eq0). + auto;progress. + + exact size_ge0. + + by rewrite take0;have[]_[]->:=H. + + by rewrite drop0. + + split;case:H=>//=;smt(size_ge0). + + smt(size_ge0 size_eq0). + + smt(size_ge0 size_eq0). + + smt(). + case(0 < n{1});last by rcondf{1}1;2:rcondf{2}1;auto;progress. + splitwhile{1} 1 : i + 1 < n;splitwhile{2} 1 : i + 1 < n. + rcondt{1}2;2:rcondt{2}2;auto;progress. + + by while(i ={i,n,glob P,C.c} /\ lres{1} = z0{2} /\ b{1} = sa{2} + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ Redo.prefixes{1}.[format p{1} (i{1}+1)] = Some (sa,sc){2});progress. + while(={i,n,glob P,C.c} /\ lres{1} = z0{2} /\ b{1} = sa{2} /\ 0 <= i{1} < n{1} + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1} /\ valid p{1} + /\ Redo.prefixes{1}.[format p{1} (i{1}+1)] = Some (sa,sc){2});last first. + + auto;1:smt(nseq0 cats0). + sp;if;auto;sp. + splitwhile{1}1: i1 < size p1 - 1. + rcondt{1}2;2:rcondf{1}4;1,2:auto. + + while(i1 < size p1);auto;2:smt(size_cat size_nseq size_ge0 size_eq0 valid_spec). + by if;auto;1:smt();sp;if;auto;progress;smt(). + + seq 1 : (i1 = size p1 - 1). + - while(i1 < size p1);auto;2:smt(size_cat size_nseq size_ge0 size_eq0 valid_spec). + by if;auto;1:smt();sp;if;auto;progress;smt(). + if;auto;sp;if;auto;smt(). + seq 1 0 : (={i,n,glob P,C.c} /\ x0{2} = (sa{2}, sc{2}) /\ 0 < i{1} < n{1} + /\ p1{1} = format p{1} (i{1} + 1) /\ (sa0,sc0){1} = x0{2} + /\ i1{1} = size p{1} + i{1} - 1 /\ lres{1} = z0{2} /\ valid p{1} + /\ Redo.prefixes{1}.[format p{1} i{1}] = Some (sa{2}, sc{2}) + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1} + /\ valid p{1});last first. + + if{1};auto. + + rcondf{2}1;auto;progress. + + move:H5;rewrite take_oversize;1:rewrite size_cat size_nseq ler_maxr/#. + move=>H_dom;rewrite domE. + by have<-:=lemma4 _ _ _ _ _ _ _ _ _ H3 H H2 H_dom;rewrite-domE. + + move:H5;rewrite take_oversize;1:rewrite size_cat size_nseq ler_maxr/#;move=>H_dom. + by have:=lemma4 _ _ _ _ _ _ _ _ _ H3 H H2 H_dom;smt(domE). + + smt(). + + move:H5;rewrite take_oversize;1:rewrite size_cat size_nseq ler_maxr/#;move=>H_dom. + by have:=lemma4 _ _ _ _ _ _ _ _ _ H3 H H2 H_dom;smt(domE). + sp;if;auto;progress. + + move:H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite Block.WRing.addr0. + + move:H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite Block.WRing.addr0. + + move:H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite Block.WRing.addr0. + + move:H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite Block.WRing.addr0. + + move:H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite Block.WRing.addr0. + + smt(). + + move:H5 H6;rewrite nth_cat nth_nseq;1:smt(size_ge0). + have->/=:!size p{1} + i{2} - 1 < size p{1} by smt(). + rewrite Block.WRing.addr0 !get_setE/= take_oversize;1:rewrite size_cat size_nseq/#. + move=>H_dom_iS H_dom_p. + have:=lemma2' 0 C.c{2} Perm.m{2}.[(sa{2}, sc{2}) <- y0L] + Perm.mi{2}.[y0L <- (sa{2}, sc{2})] Redo.prefixes{1} + p{1} (i{2}+1) sa{2} sc{2} _ _ H4 _ H_dom_iS. + + by rewrite INV_Real_addm_mi//=;smt(supp_dexcepted). + + smt(). + + by rewrite mem_set. + by rewrite!get_setE/=H2/=;smt(). + + by rewrite!get_setE/=take_oversize//=size_cat size_nseq/#. + + rewrite nth_cat;have->/=:! size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite nth_nseq//=1:/# Block.WRing.addr0. + + smt(). + + move:H5 H6;rewrite take_oversize 1:size_cat 1:size_nseq 1:/#. + rewrite nth_cat;have->/=:! size p{1} + i{2} - 1 < size p{1} by smt(). + rewrite nth_nseq//=1:/# Block.WRing.addr0 =>h1 h2. + by have:=lemma2' 0 C.c{2} Perm.m{2} Perm.mi{2} Redo.prefixes{1} + p{1} (i{2}+1) sa{2} sc{2} H3 _ H1 h2 h1;smt(). + + move:H5 H6;rewrite take_oversize 1:size_cat 1:size_nseq 1:/#. + rewrite nth_cat;have->/=:! size p{1} + i{2} - 1 < size p{1} by smt(). + by rewrite nth_nseq//=1:/# Block.WRing.addr0 !get_setE//=. + alias{1} 1 pref = Redo.prefixes;sp. + conseq(:_==> ={glob P} + /\ p1{1} = format p{1} (i{1} + 1) /\ pref{1} = Redo.prefixes{1} + /\ i1{1} = size p1{1} - 1 + /\ Redo.prefixes{1}.[take i1{1} p1{1}] = Some (sa0{1}, sc0{1}) + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1});progress. + + smt(). + + move:H9;rewrite take_format/=1:/#;1:smt(size_ge0 size_cat size_nseq). + pose x := if _ then _ else _ ;have->/={x}: x = format p{1} (i_R+1). + + rewrite/x size_cat size_nseq/=!ler_maxr 1:/#-(addzA _ _ (-1))-(addzA _ _ (-1))/=. + case(size p{1} + i_R <= size p{1})=>//=h;2:smt(size_ge0 size_cat size_nseq). + have->>/=:i_R = 0 by smt(). + by rewrite take_size/format nseq0 cats0. + by rewrite H3/==>[][]->>->>. + + move:H9;rewrite take_format/=1:/#;1:smt(size_ge0 size_cat size_nseq). + pose x := if _ then _ else _ ;have->/={x}: x = format p{1} (i_R+1). + + rewrite/x size_cat size_nseq/=!ler_maxr 1:/#-(addzA _ _ (-1))-(addzA _ _ (-1))/=. + case(size p{1} + i_R <= size p{1})=>//=h;2:smt(size_ge0 size_cat size_nseq). + have->>/=:i_R = 0 by smt(). + by rewrite take_size/format nseq0 cats0. + by rewrite H3/=. + + by rewrite size_cat size_nseq;smt(). + while{1}(={glob P} /\ 0 <= i1{1} <= size p1{1} - 1 /\ 0 < i{1} < n{1} + /\ p1{1} = format p{1} (i{1} + 1) /\ pref{1} = Redo.prefixes{1} + /\ format p{1} i{1} \in pref{1} + /\ Redo.prefixes{1}.[take i1{1} p1{1}] = Some (sa0{1}, sc0{1}) + /\ INV_Real 0 C.c{1} Perm.m{1} Perm.mi{1} Redo.prefixes{1}) + (size p1{1}-i1{1}-1);auto;last first. + + progress. + + smt(size_cat size_nseq size_ge0 size_eq0 valid_spec). + + smt(). + + by rewrite domE H3. + + by rewrite take0;have[]_[]:=H1. + + smt(). + + smt(). + rcondt 1;auto;progress. + + have->:take (i1{hr} + 1) (format p{hr} (i{hr} + 1)) = + take (i1{hr} + 1) (format p{hr} i{hr});2:smt(all_prefixes_of_INV_real domE). + rewrite!take_format;smt(valid_spec size_ge0 size_eq0 size_cat size_nseq). + + smt(). + + smt(valid_spec size_ge0 size_eq0 size_cat size_nseq). + + have->:take (i1{hr} + 1) (format p{hr} (i{hr} + 1)) = + take (i1{hr} + 1) (format p{hr} i{hr});2:smt(all_prefixes_of_INV_real domE). + rewrite!take_format;smt(valid_spec size_ge0 size_eq0 size_cat size_nseq). + smt(). + qed. + + + + lemma pr_real (D <: DISTINGUISHER{-SLCommon.C, -C, -Perm, -Redo}) &m : + Pr [ GReal(A(D)).main() @ &m : res /\ SLCommon.C.c <= max_size] = + Pr [ RealIndif(Sponge,P,DRestr(D)).main() @ &m : res]. + proof. + have->:Pr [ RealIndif(Sponge, P, DRestr(D)).main() @ &m : res ] = + Pr [ NIndif(Squeeze(SqueezelessSponge(P)),P,DRestr(D)).main() @ &m : res /\ C.c <= max_size ]. + + by rewrite eq_sym;byequiv (squeeze_squeezeless D)=>//=. + byequiv (equiv_sponge D)=>//=;progress;smt(). + qed. + +end section Real. + + +section Real_Ideal. + (* REAL & IDEAL *) + declare module D <: DISTINGUISHER {-SLCommon.C, -C, -Perm, -Redo, -F.RO, -F.FRO, -S, -BIRO.IRO, -BIRO2.IRO, -F2.RO, -F2.FRO}. + + declare axiom D_lossless (F0 <: DFUNCTIONALITY{-D}) (P0 <: DPRIMITIVE{-D}) : + islossless P0.f => islossless P0.fi => islossless F0.f => + islossless D(F0, P0).distinguish. + + + lemma A_lossless (F <: SLCommon.DFUNCTIONALITY{-A(D)}) + (P0 <: SLCommon.DPRIMITIVE{-A(D)}) : + islossless P0.f => + islossless P0.fi => islossless F.f => islossless A(D, F, P0).distinguish. + proof. + progress;proc;inline*;sp;wp. + call(:true);auto. + + exact D_lossless. + + proc;inline*;sp;if;auto;call H;auto. + + proc;inline*;sp;if;auto;call H0;auto. + proc;inline*;sp;if;auto;sp;if;auto. + while(true)(n-i);auto. + + by sp;if;auto;1:call H1;auto;smt(). + call H1;auto;smt(). + qed. + + lemma concl &m : + Pr [ RealIndif(Sponge,P,DRestr(D)).main() @ &m : res ] <= + Pr [ IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res ] + + (max_size ^ 2 - max_size)%r / 2%r / (2^r)%r / (2^c)%r + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + rewrite-(pr_real D &m). + rewrite-(equiv_ideal D &m). + have:=Real_Ideal (A(D)) A_lossless &m. + pose x:=witness;elim:x=>a b. + rewrite/dstate dprod1E DBlock.dunifin1E DCapacity.dunifin1E/= + block_card capacity_card;smt(). + qed. + + +end section Real_Ideal. + + +require import AdvAbsVal. + +section Real_Ideal_Abs. + + declare module D <: DISTINGUISHER {-SLCommon.C, -C, -Perm, -Redo, -F.RO, -F.FRO, -S, -BIRO.IRO, -BIRO2.IRO, -F2.RO, -F2.FRO}. + + declare axiom D_lossless (F0 <: DFUNCTIONALITY{-D}) (P0 <: DPRIMITIVE{-D}) : + islossless P0.f => islossless P0.fi => islossless F0.f => + islossless D(F0, P0).distinguish. + + + local module Neg_D (D : DISTINGUISHER) (F : DFUNCTIONALITY) (P : DPRIMITIVE) = { + proc distinguish () : bool = { + var b : bool; + b <@ D(F,P).distinguish(); + return !b; + } + }. + + + local lemma Neg_D_lossless (F <: DFUNCTIONALITY{-Neg_D(D)}) (P <: DPRIMITIVE{-Neg_D(D)}) : + islossless P.f => islossless P.fi => + islossless F.f => islossless Neg_D(D, F, P).distinguish. + proof. + by progress;proc;inline*;call(D_lossless F P H H0 H1);auto. + qed. + + + local lemma useful m mi a : + invm m mi => ! a \in m => Distr.is_lossless ((bdistr `*` cdistr) \ rng m). + proof. + move=>hinvm nin_dom. + have prod_ll:Distr.is_lossless (bdistr `*` cdistr). + + by rewrite dprod_ll DBlock.dunifin_ll DCapacity.dunifin_ll. + apply dexcepted_ll=>//=;rewrite-prod_ll. + have->:predT = predU (predC (rng m)) (rng m);1:rewrite predCU//=. + rewrite Distr.mu_disjoint 1:predCI//= RField.addrC. + have/=->:=ltr_add2l (mu (bdistr `*` cdistr) (rng m)) 0%r. + rewrite Distr.witness_support/predC. + move:nin_dom;apply absurd=>//=;rewrite negb_exists/==>hyp. + have{hyp}hyp:forall x, rng m x by smt(supp_dprod DBlock.supp_dunifin DCapacity.supp_dunifin). + move:a. + have:=eqEcard (fdom m) (frng m);rewrite leq_card_rng_dom/=. + have->//=:fdom m \subset frng m. + + by move=> x; rewrite mem_fdom mem_frng hyp. + smt(mem_fdom mem_frng). + qed. + + local lemma invmC (m mi : (state, state) fmap) : + invm m mi <=> invm mi m. + proof. smt(). qed. + + + local lemma Real_lossless : + islossless RealIndif(Sponge, P, DRestr(Neg_D(D))).main. + proof. + proc;inline*;auto;call(: invm Perm.m Perm.mi);2..:auto. + + exact D_lossless. + + proc;inline*;sp;if;auto;sp;if;auto;progress. + - by have:=useful _ _ _ H H1. + - smt(invm_set dexcepted1E). + + proc;inline*;sp;if;auto;sp;if;auto;progress. + - have:=H;rewrite invmC=>h;have/#:=useful _ _ _ h H1. + - move:H;rewrite invmC=>H;rewrite invmC;smt(invm_set dexcepted1E domE rngE). + + proc;inline*;sp;if;auto;sp;if;auto. + while(invm Perm.m Perm.mi)(n-i);auto. + - sp;if;auto;2:smt();sp;if;auto;2:smt();progress. + * by have:=useful _ _ _ H H2. + * smt(invm_set dexcepted1E). + smt(). + conseq(:_==> invm Perm.m Perm.mi);1:smt(). + while(invm Perm.m Perm.mi)(size xs);auto. + - sp;if;auto;progress. + * by have:=useful _ _ _ H H1. + * smt(invm_set dexcepted1E). + * smt(size_behead). + * smt(size_behead). + smt(size_ge0 size_eq0). + smt(emptyE). + qed. + + + local lemma Ideal_lossless : + islossless IdealIndif(BIRO.IRO, SimLast(S), DRestr(Neg_D(D))).main. + proof. + proc;inline*;auto;call(D_lossless (FC(BIRO.IRO)) (PC(SimLast(S, BIRO.IRO))) _ _ _);auto. + + proc;inline*;sp;if;auto;sp;if;auto;sp;if;auto;2:smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + sp;if;auto;sp;if;auto;2,4:smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + * while(true)(n-i);auto;2:smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + by sp;if;auto;smt(DBlock.dunifin_ll). + while(true)(n0-i0);auto;2:smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + by sp;if;auto;smt(DBlock.dunifin_ll). + + by proc;inline*;sp;if;auto;sp;if;auto;smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + proc;inline*;sp;if;auto;sp;if;auto;while(true)(n-i);auto;2:smt(). + by sp;if;auto;smt(DBlock.dunifin_ll). + qed. + + + + + local lemma neg_D_concl &m : + Pr [ IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res ] <= + Pr [ RealIndif(Sponge,P,DRestr(D)).main() @ &m : res ] + + (max_size ^ 2 - max_size)%r / 2%r / (2^r)%r / (2^c)%r + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + have->:Pr[IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res] = + Pr[Neg_main(IdealIndif(BIRO.IRO, SimLast(S), DRestr(Neg_D(D)))).main() @ &m : res]. + + by byequiv=>//=;proc;inline*;auto;conseq(:_==> b0{1} = b2{2});progress;sim. + have->:Pr [ RealIndif(Sponge,P,DRestr(D)).main() @ &m : res ] = + Pr [ Neg_main(RealIndif(Sponge,P,DRestr(Neg_D(D)))).main() @ &m : res ]. + + by byequiv=>//=;proc;inline*;auto;conseq(:_==> b0{1} = b2{2});progress;sim. + have h1 := Neg_A_Pr_minus (RealIndif(Sponge,P,DRestr(Neg_D(D)))) &m Real_lossless. + have h2 := Neg_A_Pr_minus (IdealIndif(BIRO.IRO, SimLast(S), DRestr(Neg_D(D)))) &m Ideal_lossless. + have/#:=concl (Neg_D(D)) _ &m;progress. + by proc;call(D_lossless F0 P0 H H0 H1);auto. + qed. + + lemma Inefficient_Real_Ideal &m : + `|Pr [ RealIndif(Sponge,Perm,DRestr(D)).main() @ &m : res ] - + Pr [ IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res ]| <= + (max_size ^ 2 - max_size)%r / 2%r / (2^r)%r / (2^c)%r + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + have := concl D D_lossless &m. + have := neg_D_concl &m. + pose p1 := Pr[IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res]. + pose p2 := Pr[RealIndif(Sponge, Perm, DRestr(D)).main() @ &m : res]. + rewrite-5!(RField.addrA). + pose p3 := (max_size ^ 2)%r / 2%r / (2 ^ r)%r / (2 ^ c)%r + + (max_size%r * ((2 * max_size)%r / (2 ^ c)%r) + + max_size%r * ((2 * max_size)%r / (2 ^ c)%r)). + smt(). + qed. + +end section Real_Ideal_Abs. + + + +module Simulator (F : DFUNCTIONALITY) = { + var m : (state, state) fmap + var mi : (state, state) fmap + var paths : (capacity, block list * block) fmap + var unvalid_map : (block list * int, block) fmap + proc init() = { + m <- empty; + mi <- empty; + paths <- empty.[c0 <- ([],b0)]; + unvalid_map <- empty; + } + proc f (x : state) : state = { + var p,v,q,k,cs,y,y1,y2,r; + if (x \notin m) { + if (x.`2 \in paths) { + (p,v) <- oget paths.[x.`2]; + (q,k) <- parse (rcons p (v +^ x.`1)); + if (valid q) { + cs <@ F.f(q, k); + y1 <- last b0 cs; + } else { + if (0 < k) { + if ((q,k-1) \notin unvalid_map) { + r <$ bdistr; + unvalid_map.[(q,k-1)] <- r; + } + y1 <- oget unvalid_map.[(q,k-1)]; + } else { + y1 <- b0; + } + } + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + m.[x] <- y; + mi.[y] <- x; + if (x.`2 \in paths) { + (p,v) <-oget paths.[x.`2]; + paths.[y2] <- (rcons p (v +^ x.`1),y.`1); + } + } else { + y <- oget m.[x]; + } + return y; + } + proc fi (x : state) : state = { + var y,y1,y2; + if (! x \in mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + mi.[x] <- y; + m.[y] <- x; + } else { + y <- oget mi.[x]; + } + return y; + } +}. + +section Simplify_Simulator. + +declare module D <: DISTINGUISHER {-Simulator, -F.RO, -BIRO.IRO, -C, -S, -BIRO2.IRO}. + +declare axiom D_lossless (F0 <: DFUNCTIONALITY{-D}) (P0 <: DPRIMITIVE{-D}) : + islossless P0.f => islossless P0.fi => islossless F0.f => + islossless D(F0, P0).distinguish. + +local clone import PROM.FullRO as IRO2 with + type in_t <- block list * int, + type out_t <- block, + op dout _ <- bdistr, + type d_in_t <- unit, + type d_out_t <- bool. +import FullEager. + +local module Simu (FRO : IRO2.RO) (F : DFUNCTIONALITY) = { + proc init() = { + Simulator(F).init(); + FRO.init(); + } + proc f (x : state) : state = { + var p,q,v,k,i,cs,y,y1,y2; + if (x \notin Simulator.m) { + if (x.`2 \in Simulator.paths) { + (p,v) <- oget Simulator.paths.[x.`2]; + (q,k) <- parse (rcons p (v +^ x.`1)); + if (valid q) { + cs <@ F.f(q, k); + y1 <- last b0 cs; + } else { + if (0 < k) { + i <- 0; + while (i < k) { + FRO.sample(q,i); + i <- i + 1; + } + y1 <@ FRO.get(q,k-1); + } else { + y1 <- b0; + } + } + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1,y2); + Simulator.m.[x] <- y; + Simulator.mi.[y] <- x; + if (x.`2 \in Simulator.paths) { + (p,v) <-oget Simulator.paths.[x.`2]; + Simulator.paths.[y2] <- (rcons p (v +^ x.`1),y.`1); + } + } else { + y <- oget Simulator.m.[x]; + } + return y; + } + proc fi (x : state) : state = { + var y,y1,y2; + if (! x \in Simulator.mi) { + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + Simulator.mi.[x] <- y; + Simulator.m.[y] <- x; + } else { + y <- oget Simulator.mi.[x]; + } + return y; + } +}. + +local module L (F : IRO2.RO) = { + proc distinguish = IdealIndif(BIRO.IRO, Simu(F), DRestr(D)).main +}. + +local lemma equal1 &m : + Pr [ IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res ] = + Pr [ L(IRO2.RO).distinguish() @ &m : res ]. +proof. +byequiv=>//=; proc; inline*; auto. +call (: ={BIRO.IRO.mp,C.c} /\ ={m,mi,paths}(S,Simulator) /\ + BIRO2.IRO.mp{1} = IRO2.RO.m{2} /\ + incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1}); first last. ++ by proc; inline*; conseq=>/>; sim. ++ by proc; inline*; conseq=>/>; sim. ++ by auto. +proc; sp; if; auto. +call(: ={BIRO.IRO.mp} /\ ={m,mi,paths}(S,Simulator) /\ + BIRO2.IRO.mp{1} = IRO2.RO.m{2} /\ + incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1});auto. +if; 1,3: by auto. +seq 1 1: (={BIRO.IRO.mp,y1,x} /\ ={m,mi,paths}(S,Simulator) /\ + BIRO2.IRO.mp{1} = IRO2.RO.m{2} /\ + incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1}); last first. +- by conseq=>/>; auto. +if; 1,3: by auto. +inline*; sp; if; 1,2: auto. +- move=> /> &1 &2 h1 h2 bl n h3 h4 h5 h6 h7 h8. + have:= h1; rewrite-h3 /= => [#] ->>->>. + have:= h4; rewrite-h2 /= => [#] ->>->>. + have->>/=: q{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`1 by smt(). + have->>/=: k{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`2 by smt(). + move: h5; have-> h5:= formatK (rcons p{1} (v{1} +^ x{2}.`1)). + by have->>/=: q{1} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`1 by smt(). +- sp; if; auto. + * move=> /> &1 &2 h1 h2 bl n h3 h4 h5 h6 h7 h8 h9 h10. + have:= h1; rewrite-h3 /= => [#] ->>->>. + have:= h4; rewrite-h2 /= => [#] ->>->>. + have->>/=: q{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`1 by smt(). + have->>/=: k{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`2 by smt(). + move: h5; have-> h5:= formatK (rcons p{1} (v{1} +^ x{2}.`1)). + by have->>/=: q{1} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`1 by smt(). + by conseq(:_ ==> ={bs, BIRO.IRO.mp})=> />; sim=> />; smt(parseK formatK). +sp; rcondt{1} 1; 1: auto; if{2}; last first. ++ by rcondf{1} 1; auto; smt(parseK formatK). +sp; rcondf{2} 4; 1: auto. ++ conseq(:_ ==> (q,k-1) \in RO.m)=> />. + splitwhile 1 : i + 1 < k. + rcondt 2; 1:(auto; while (i + 1 <= k); auto; smt()). + rcondf 7; 1:(auto; while (i + 1 <= k); auto; smt()). + seq 1 : (i + 1 = k). + - by while(i + 1 <= k); auto; smt(). + by auto=> />; smt(mem_set). +wp; rnd{2}; wp=> /=; conseq=> />. +conseq(:_==> i{2} = k{2} /\ + (0 < i{2} => last Block.b0 bs0{1} = oget RO.m{2}.[(q{2}, i{2} -1)]) /\ + BIRO2.IRO.mp{1} = RO.m{2} /\ incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1}) =>/>. ++ smt(DBlock.dunifin_ll). +while (i{2} <= k{2} /\ n0{1} = k{2} /\ i0{1} = i{2} /\ x1{1} = q{2} /\ ={k} /\ + (0 < i{2} => last Block.b0 bs0{1} = oget RO.m{2}.[(q{2}, i{2} - 1)]) /\ + BIRO2.IRO.mp{1} = RO.m{2} /\ incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1}). ++ sp; wp 2 2=> /=; conseq=> />. + conseq(:_==> b0{1} = oget RO.m{2}.[(q{2}, i{2})] /\ + BIRO2.IRO.mp{1} = RO.m{2} /\ + incl Simulator.unvalid_map{2} BIRO2.IRO.mp{1}); 1: smt(last_rcons). + if{1}; 2: rcondf{2} 2; 1: rcondt{2} 2; 1,3: auto. + - by auto=> />; smt(incl_upd_nin). + by auto; smt(DBlock.dunifin_ll). +auto=> /> &1 &2 h1 h2 [#] q_L k_L h3 h4 h5 h6 h7 h8 h9 h10;split. ++ have:= h1; rewrite -h3 => [#] />; have:= h4; rewrite -h2 => [#] />. + have:= h5. + have-> : q{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`1 by smt(). + have-> : k{2} = (parse (rcons p{1} (v{1} +^ x{2}.`1))).`2 by smt(). + by rewrite (formatK (rcons p{1} (v{1} +^ x{2}.`1)))=> [#] />; smt(). +smt(). +qed. + + +local lemma equal2 &m : + Pr [ IdealIndif(BIRO.IRO, Simulator, DRestr(D)).main() @ &m : res ] = + Pr [ L(IRO2.LRO).distinguish() @ &m : res ]. +proof. +byequiv=>//=; proc; inline*; auto. +call (: ={BIRO.IRO.mp,C.c,Simulator.m,Simulator.mi,Simulator.paths} /\ + Simulator.unvalid_map{1} = IRO2.RO.m{2}); first last. ++ by proc; inline*; conseq=> />; sim. ++ by proc; inline*; conseq=> />; sim. ++ by auto=> />. +proc; sp; if; auto. +call(: ={BIRO.IRO.mp,Simulator.m,Simulator.mi,Simulator.paths} /\ + Simulator.unvalid_map{1} = IRO2.RO.m{2}); auto. +if; 1,3: auto. +seq 1 1: (={y1,x, BIRO.IRO.mp, Simulator.m, Simulator.mi, Simulator.paths} /\ + Simulator.unvalid_map{1} = RO.m{2}); 2: by (conseq=> />; sim). +if; 1,3: auto; sp. +conseq=> />. +conseq(: ={q, k, BIRO.IRO.mp} /\ Simulator.unvalid_map{1} = RO.m{2} ==> _)=> />. ++ move=> &1 &2 h1 h2 h3 h4 h5 h6. + by have:= h1; rewrite -h3 => [#] /> /#. +inline*; if; 1: auto; 1: sim. +if; 1,3: auto; sp. +swap{2} 4; while{2}((i<=k){2})(k{2}-i{2}); 1: by (auto; smt()). +by sp; if{1}; 2: rcondf{2} 2; 1: rcondt{2} 2; auto; smt(DBlock.dunifin_ll). +qed. + + + +lemma Simplify_simulator &m : + Pr [ IdealIndif(BIRO.IRO, Simulator, DRestr(D)).main() @ &m : res ] = + Pr [ IdealIndif(BIRO.IRO, SimLast(S), DRestr(D)).main() @ &m : res ]. +proof. +rewrite (equal1 &m) (equal2 &m) eq_sym. +by byequiv(RO_LRO_D L _)=> //=; exact/dunifin_ll. +qed. + + +end section Simplify_Simulator. + + + + + +section Real_Ideal. + declare module D <: DISTINGUISHER {-SLCommon.C, -C, -Perm, -Redo, -F.RO, -F.FRO, -S, -BIRO.IRO, -BIRO2.IRO, -F2.RO, -F2.FRO, -Simulator}. + + declare axiom D_lossless (F0 <: DFUNCTIONALITY{-D}) (P0 <: DPRIMITIVE{-D}) : + islossless P0.f => islossless P0.fi => islossless F0.f => + islossless D(F0, P0).distinguish. + + + lemma Real_Ideal &m : + `|Pr [ RealIndif(Sponge,Perm,DRestr(D)).main() @ &m : res ] - + Pr [ IdealIndif(BIRO.IRO, Simulator, DRestr(D)).main() @ &m : res ]| <= + (max_size ^ 2 - max_size)%r / 2%r / (2^r)%r / (2^c)%r + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + rewrite(Simplify_simulator D D_lossless &m). + exact/(Inefficient_Real_Ideal D D_lossless &m). + qed. + +end section Real_Ideal. diff --git a/sha3/proof/smart_counter/Gext.eca b/sha3/proof/smart_counter/Gext.eca new file mode 100644 index 0000000..e90c3a1 --- /dev/null +++ b/sha3/proof/smart_counter/Gext.eca @@ -0,0 +1,726 @@ +pragma -oldip. +require import Core Int Real StdOrder Ring StdBigop. +require import List FSet SmtMap Common SLCommon FelTactic Mu_mem. +require import DProd Dexcepted PROM. +(*...*) import Capacity IntID IntOrder Bigreal RealOrder BRA DCapacity. + +require (*..*) Gcol. + +clone export Gcol as Gcol0. + +op bad_ext (m mi:smap) y = + mem (image snd (fdom m)) y \/ + mem (image snd (fdom mi)) y. + +op hinvc (m:(handle,capacity)fmap) (c:capacity) = + find (+ pred1 c) m. + +module G2(D:DISTINGUISHER,HS:FRO) = { + + module M = { + + proc f(p : block list): block = { + var sa, sa'; + var h, i, counter <- 0; + sa <- b0; + while (i < size p ) { + if ((sa +^ nth witness p i, h) \in G1.mh) { + (sa, h) <- oget G1.mh.[(sa +^ nth witness p i, h)]; + } else { + if (counter < size p - prefix p (get_max_prefix p (elems (fdom C.queries)))) { + HS.sample(G1.chandle); + sa' <@ F.RO.get(take (i+1) p); + sa <- sa +^ nth witness p i; + G1.mh.[(sa,h)] <- (sa', G1.chandle); + G1.mhi.[(sa',G1.chandle)] <- (sa, h); + (sa,h) <- (sa',G1.chandle); + G1.chandle <- G1.chandle + 1; + counter <- counter + 1; + } + } + i <- i + 1; + } + sa <@ F.RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2, hy2, hx2, handles_,t; + + if (x \notin G1.m) { + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + y1 <@ F.RO.get (rcons p (v +^ x.`1)); + y2 <$ cdistr; + } else { + y1 <$ bdistr; + y2 <$ cdistr; + } + y <- (y1, y2); + + handles_ <@ HS.allKnown(); + if (!rng handles_ x.`2) { + HS.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ HS.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + t <@ HS.queried((oget G1.mh.[(x.`1,hx2)]).`2, Unknown); + if ((x.`1, hx2) \in G1.mh /\ t) { + hy2 <- (oget G1.mh.[(x.`1, hx2)]).`2; + y2 <@ HS.get(hy2); + G1.bext <- G1.bext \/ bad_ext G1.m G1.mi y2 \/ y2 = x.`2; + y <- (y.`1, y2); + G1.m.[x] <- y; + G1.mi.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + HS.set(hy2, y.`2); + G1.m.[x] <- y; + G1.mh.[(x.`1, hx2)] <- (y.`1, hy2); + G1.mi.[y] <- x; + G1.mhi.[(y.`1, hy2)] <- (x.`1, hx2); + } + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + G1.paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget G1.m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2, hx2, hy2, handles_, t; + + if (x \notin G1.mi) { + handles_ <@ HS.allKnown(); + if (!rng handles_ x.`2) { + HS.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ HS.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + t <@ HS.queried((oget G1.mhi.[(x.`1,hx2)]).`2, Unknown); + if ((x.`1, hx2) \in G1.mhi /\ t) { + (y1,hy2) <- oget G1.mhi.[(x.`1, hx2)]; + y2 <@ HS.get(hy2); + y <- (y.`1, y2); + G1.bext <- G1.bext \/ bad_ext G1.m G1.mi y2 \/ y2 = x.`2; + G1.mi.[x] <- y; + G1.m.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + HS.set(hy2, y.`2); + G1.mi.[x] <- y; + G1.mhi.[(x.`1, hx2)] <- (y.`1, hy2); + G1.m.[y] <- x; + G1.mh.[(y.`1, hy2)] <- (x.`1, hx2); + } + } else { + y <- oget G1.mi.[x]; + } + return y; + } + + } + + proc distinguish(): bool = { + var b; + + F.RO.m <- empty; + G1.m <- empty; + G1.mi <- empty; + G1.mh <- empty; + G1.mhi <- empty; + G1.bext <- false; + C.queries<- empty.[[] <- b0]; + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + HS.set(0,c0); + G1.paths <- empty.[c0 <- ([<:block>],b0)]; + G1.chandle <- 1; + b <@ D(M,S).distinguish(); + return b; + } +}. + +clone include EagerCore +proof * by (move=> _; exact/Capacity.DCapacity.dunifin_ll). + +section. + declare module D <: DISTINGUISHER{-G1, -G2, -FRO, -C}. + + op inv_ext (m mi:smap) (FROm:handles) = + exists x h, mem (fdom m `|` fdom mi) x /\ FROm.[h] = Some (x.`2, Unknown). + + op inv_ext1 bext1 bext2 (m mi:smap) (FROm:handles) = + bext1 => (bext2 \/ inv_ext m mi FROm). + + lemma rng_restr (m : ('from, 'to * 'flag) fmap) f x: + rng (restr f m) x <=> rng m (x,f). + proof. + rewrite !rngE;split=>-[z]H;exists z;move:H;rewrite restrP; case m.[z]=>//=. + by move=> [t f'] /=;case (f'=f). + qed. + + equiv G1_G2 : G1(DRestr(D)).main ~ Eager(G2(DRestr(D))).main1 : + ={glob D} ==> ={res} /\ inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2}. + proof. + proc;inline{2} FRO.init G2(D, FRO).distinguish;wp. + inline*;wp. + call (_: ={F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.paths,G1.chandle,FRO.m,C.queries,C.c} /\ + inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2} /\ + (forall h, h \in FRO.m => h < G1.chandle){1}). + + proc. + sp;if;auto;inline G1(DRestr(D)).S.f G2(DRestr(D), FRO).S.f;sp;wp. + if=>//;last by auto. + seq 2 2: (={F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.paths,G1.chandle,FRO.m,x,x0,y0,C.queries,C.c} /\ + inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2} /\ x{1} = x0{1} /\ + (forall h, h \in FRO.m => h < G1.chandle){1} /\ + x0{1} \notin G1.m{1}). + + by if=>//;auto;call (_: ={F.RO.m});[sim |auto]. + seq 3 5: + (={F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.paths,G1.chandle,FRO.m,hx2,x,x0,y0,hx2,C.queries,C.c} /\ + t{2} = (in_dom_with FRO.m (oget G1.mh.[(x.`1, hx2)]).`2 Unknown){1} /\ x{1} = x0{1} /\ + (G1.bext{1} => (G1.bext{2} \/ (rng FRO.m (x.`2, Unknown)){2} \/ + inv_ext G1.m{2} G1.mi{2} FRO.m{2})) /\ + (forall h, h \in FRO.m => h < G1.chandle){1} /\ + x0{1} \notin G1.m{1}). + + inline *; auto=> |> &ml &mr Hi Hhand. + rewrite -dom_restr rng_restr !rngE /= negb_exists=> /= |> x_notin_G1. + case: (x0{mr})=> b0 c0 /=; split=> [x0K_notinrng_m|h mh]. + + split=> [|/#]. + split=> [[/Hi [-> //|[] x h /= [] H1 H2]|[] h Hh]|h]. + + by right; right; exists x h; rewrite get_setE /#. + + right; left; exists h; rewrite get_set_neqE //=. + by have:= Hhand h; rewrite domE Hh /#. + by rewrite mem_set /#. + split=> [/#|[/Hi [->//|[] x h' /= [] H1 H2]|[] h' Hh']]. + + by right; right; exists x h'. + by right; left; exists h'. + seq 1 1: (={x0,y0,x,F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.paths,G1.chandle,FRO.m,C.queries,C.c} /\ + inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2} /\ x{1} = x0{1} /\ + forall (h : handle), h \in FRO.m{1} => h < G1.chandle{1});2:by auto. + if=>//. + + inline *;rcondt{2} 4. + + by move=> &m;auto;rewrite /in_dom_with. +(* auto=> |>. (* Bug ???? *) *) + auto;progress. + + rewrite /inv_ext1=>/H{H}[->//|[|[[x1 x2] h [Hx Hh]]]]. + + rewrite rngE/==>[][]h Hh. + case (h = (oget G1.mh{2}.[(x0{2}.`1, hx2{2})]).`2)=> [->>|Hneq]. + + by rewrite Hh oget_some/#. + by right;exists x0{2} h; rewrite fdom_set !in_fsetU !in_fset1/= get_set_neqE//. + case (h = (oget G1.mh{2}.[(x0{2}.`1, hx2{2})]).`2)=> [->>|Hneq]. + + rewrite Hh /bad_ext oget_some /= <@ Hx;rewrite !inE. + by move=>[|]/(mem_image snd)->. + right;exists (x1,x2) h; move:Hx. + by rewrite !fdom_set !in_fsetU !in_fset1 //= => [][] -> //=; rewrite get_set_neqE. + by move:H5 H2;rewrite /in_dom_with mem_set /#. + inline *;auto;progress;last by move:H3;rewrite mem_set /#. + rewrite /inv_ext1=> /H [->//|[|[x' h [Hx Hh]]]]. + + rewrite rngE=> [][] h Hh. + right;exists x0{2} h; rewrite fdom_set !inE /= get_set_neqE //. + by move:(H0 h);rewrite domE Hh /#. + right;exists x' h; rewrite fdom_set !inE /= !mem_fdom. + move:(H0 h);rewrite domE Hh //= !get_setE => Hh2. + have-> /= : ! h = G1.chandle{2} by smt(). + by rewrite Hh /= mem_set; move: Hx; rewrite in_fsetU !mem_fdom=>[][]->. + + + proc;sp;if;auto;inline G1(DRestr(D)).S.fi G2(DRestr(D), FRO).S.fi;sp;wp. + if=>//;last by auto. + seq 6 8: + (={F.RO.m,G1.m,G1.mi,G1.mh,G1.mhi,G1.paths,G1.chandle,FRO.m,hx2,x,x0,y0,hx2,C.queries,C.c} /\ + t{2} = (in_dom_with FRO.m (oget G1.mhi.[(x.`1, hx2)]).`2 Unknown){1} /\ x{1} = x0{1} /\ + (G1.bext{1} => (G1.bext{2} \/ (rng FRO.m (x.`2, Unknown)){2} \/ + inv_ext G1.m{2} G1.mi{2} FRO.m{2})) /\ + (forall h, h \in FRO.m => h < G1.chandle){1} /\ + x{1} \notin G1.mi{1}). + + inline *; auto=> |> &ml &mr Hi Hhand. + rewrite -dom_restr rng_restr !rngE /= negb_exists=> /= |> c_le_msize x_notin_G1. + case: (x{mr})=> b c /=; split=> [xK_notinrng_m|h mh]. + + split=> [b0 _ c0 _|/#]. + split=> [[/Hi [-> //|[] x h /= [] H1 H2]|[] h Hh]|h]. + + by right; right; exists x h; rewrite get_setE /#. + + right; left; exists h; rewrite get_set_neqE //=. + by have:= Hhand h; rewrite domE Hh /#. + by rewrite mem_set /#. + split=> [/#|b0 _ c0 _ [/Hi [->//|[] x h' /= [] H1 H2]|[] h' Hh']]. + + by right; right; exists x h'. + by right; left; exists h'. + if=>//. + + inline *;rcondt{2} 4. + + by move=> &m;auto;rewrite /in_dom_with. + auto;progress. + + rewrite /inv_ext1=>/H{H}[->//|[|[[x1 x2] h [Hx Hh]]]]. + + rewrite rngE => [][h]Hh. + case (h = (oget G1.mhi{2}.[(x0{2}.`1, hx2{2})]).`2)=> [->>|Hneq]. + + by left;rewrite Hh oget_some. + right; exists x0{2} h; rewrite !in_fsetU !mem_fdom !mem_set /=. + by rewrite get_set_neqE. + case (h = (oget G1.mhi{2}.[(x0{2}.`1, hx2{2})]).`2)=> [->>|Hneq]. + + rewrite Hh /bad_ext oget_some /= <@ Hx;rewrite !inE. + by move=>[|]/(mem_image snd)->. + right;exists (x1,x2) h;rewrite !in_fsetU !mem_fdom !mem_set /=. + rewrite get_set_neqE //= Hh /=. + by move: Hx; rewrite in_fsetU !mem_fdom=>[][] ->. + by move:H5 H1;rewrite /in_dom_with mem_set /#. + inline *;auto;progress;last by move:H3;rewrite mem_set /#. + rewrite /inv_ext1=> /H [->//|[|[x' h [Hx Hh]]]]. + + rewrite rngE => [][h]Hh. + right;exists x0{2} h;rewrite get_setE in_fsetU !mem_fdom !mem_set /=. + by move:(H0 h);rewrite domE Hh /#. + right;exists x' h;rewrite get_setE in_fsetU !mem_fdom !mem_set /=. + move:Hx; rewrite in_fsetU 2!mem_fdom=>[][]->//=. + + by move:(H0 h);rewrite domE Hh /#. + by move:(H0 h);rewrite domE Hh /#. + + + proc;sp;if;auto;sp;if;auto;sp. + inline G1(DRestr(D)).M.f G2(DRestr(D), FRO).M.f;sp;wp. + conseq (_: ={sa,G1.mh,G1.mhi,F.RO.m, G1.chandle, FRO.m,C.queries,C.c} /\ + inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2} /\ + forall (h0 : handle), h0 \in FRO.m{1} => h0 < G1.chandle{1})=>//. + sp;call (_: ={F.RO.m});1:by sim. + while (={sa,G1.mh,G1.mhi,F.RO.m,G1.chandle,FRO.m,i,h,sa,p,C.queries,counter,bs} /\ + inv_ext1 G1.bext{1} G1.bext{2} G1.m{2} G1.mi{2} FRO.m{2} /\ p{2} = bs{2} /\ + forall (h0 : handle), h0 \in FRO.m{1} => h0 < G1.chandle{1})=>//. + if=>//;inline *;1:by auto. + if;1,3:auto;progress. + rcondt{2} 3;1:by auto=>/#. + auto;progress. + + move=>bad1;have[/=->//|]:=H bad1;rewrite/inv_ext=>[][]x h[]H_dom Hh;right. + exists x h;rewrite H_dom/= get_set_neqE //=. + by move:(H0 h);rewrite domE Hh /#. + + smt(mem_set). + + move=>bad1;have[/=->//|]:=H bad1;rewrite/inv_ext=>[][]x h[]H_dom Hh;right. + exists x h;rewrite H_dom/= get_set_neqE //=. + by move:(H0 h);rewrite domE Hh /#. + + smt(mem_set). + + (* **************** *) + inline *;auto;progress. + smt(mem_set mem_empty). + qed. + +end section. + +section EXT. + + declare module D <: DISTINGUISHER{-C, -PF, -G1, -G2, -Perm, -RO, -Redo}. + + local module ReSample = { + var count:int + proc f (h:handle) = { + var c; + c <$ cdistr; + if (card (fdom G1.m) <= max_size /\ card (fdom G1.mi) <= max_size + /\ ReSample.count < max_size) { + G1.bext <- G1.bext \/ mem (image snd (fdom G1.m `|` fdom G1.mi)) c; + FRO.m.[h] <- (c,Unknown); + count <- count + 1 ; + } + } + + proc f1 (x:capacity,h:handle) = { + var c; + c <$ cdistr; + if (card (fdom G1.m) < max_size /\ card (fdom G1.mi) < max_size /\ ReSample.count < max_size) { + G1.bext <- G1.bext \/ mem (image snd (fdom G1.m `|` fdom G1.mi) `|` fset1 x) c; + FRO.m.[h] <- (c,Unknown); + count <- count + 1; + } + } + + }. + + local module Gext = { + + proc resample () = { + Iter(ReSample).iter (elems (fdom (restr Unknown FRO.m))); + } + + module M = { + + proc f(p : block list): block = { + var sa, sa'; + var h, i, counter <- 0; + sa <- b0; + while (i < size p ) { + if ((sa +^ nth witness p i, h) \in G1.mh) { + (sa, h) <- oget G1.mh.[(sa +^ nth witness p i, h)]; + } else { + if (counter < size p - prefix p (get_max_prefix p (elems (fdom C.queries)))) { + RRO.sample(G1.chandle); + sa' <@ F.RO.get(take (i+1) p); + sa <- sa +^ nth witness p i; + G1.mh.[(sa,h)] <- (sa', G1.chandle); + G1.mhi.[(sa',G1.chandle)] <- (sa, h); + (sa,h) <- (sa',G1.chandle); + G1.chandle <- G1.chandle + 1; + counter <- counter + 1; + } + } + i <- i + 1; + } + sa <@ F.RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2, hy2, hx2, handles_,t; + + if (x \notin G1.m) { + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + y1 <@ F.RO.get (rcons p (v +^ x.`1)); + } else { + y1 <$ bdistr; + } + y2 <$ cdistr; + y <- (y1, y2); + (* exists x h, mem (dom G1.m) x /\ handles.[h] = Some (x.2, I) *) + + handles_ <@ RRO.allKnown(); + if (!rng handles_ x.`2) { + RRO.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ RRO.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + t <@ RRO.queried((oget G1.mh.[(x.`1,hx2)]).`2, Unknown); + if ((x.`1, hx2) \in G1.mh /\ t) { + hy2 <- (oget G1.mh.[(x.`1, hx2)]).`2; + ReSample.f1(x.`2, hy2); + y2 <@ FRO.get(hy2); + y <- (y.`1, y2); + G1.m.[x] <- y; + G1.mi.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + RRO.set(hy2, y.`2); + G1.m.[x] <- y; + G1.mh.[(x.`1, hx2)] <- (y.`1, hy2); + G1.mi.[y] <- x; + G1.mhi.[(y.`1, hy2)] <- (x.`1, hx2); + } + if (x.`2 \in G1.paths) { + (p,v) <- oget G1.paths.[x.`2]; + G1.paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget G1.m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2, hx2, hy2, handles_, t; + + if (x \notin G1.mi) { + handles_ <@ RRO.allKnown(); + if (!rng handles_ x.`2) { + RRO.set(G1.chandle, x.`2); + G1.chandle <- G1.chandle + 1; + } + handles_ <@ RRO.allKnown(); + hx2 <- oget (hinvc handles_ x.`2); + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + t <@ RRO.queried((oget G1.mhi.[(x.`1,hx2)]).`2, Unknown); + if ((x.`1, hx2) \in G1.mhi /\ t) { + (y1,hy2) <- oget G1.mhi.[(x.`1, hx2)]; + ReSample.f1(x.`2,hy2); + y2 <@ FRO.get(hy2); + y <- (y.`1, y2); + + G1.mi.[x] <- y; + G1.m.[y] <- x; + } else { + hy2 <- G1.chandle; + G1.chandle <- G1.chandle + 1; + RRO.set(hy2, y.`2); + G1.mi.[x] <- y; + G1.mhi.[(x.`1, hx2)] <- (y.`1, hy2); + G1.m.[y] <- x; + G1.mh.[(y.`1, hy2)] <- (x.`1, hx2); + } + } else { + y <- oget G1.mi.[x]; + } + return y; + } + + } + + proc distinguish(): bool = { + var b; + + SLCommon.C.c <- 0; + F.RO.m <- empty; + G1.m <- empty; + G1.mi <- empty; + G1.mh <- empty; + G1.mhi <- empty; + G1.bext <- false; + ReSample.count <- 0; + FRO.m <- empty; + + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + RRO.set(0,c0); + G1.paths <- empty.[c0 <- ([<:block>],b0)]; + G1.chandle <- 1; + b <@ DRestr(D,M,S).distinguish(); + resample(); + return b; + } + }. + + op inv_lt (m2 mi2:smap) c1 (Fm2:handles) count2 = + card (fdom m2) < c1 /\ card (fdom mi2) < c1 /\ + count2 + card (fdom (restr Unknown Fm2)) < c1 /\ + c1 <= max_size. + + op inv_le (m2 mi2:smap) c1 (Fm2:handles) count2 = + card (fdom m2) <= c1 /\ card (fdom mi2) <= c1 /\ + count2 + card (fdom (restr Unknown Fm2)) <= c1 /\ + c1 <= max_size. + + lemma fset0_eqP (s:'a fset): s = fset0 <=> forall x, !mem s x. + proof. + split=>[-> x|Hmem];1:by rewrite inE. + by apply fsetP=>x;rewrite inE Hmem. + qed. + + lemma size_set (m:('a,'b)fmap) (x:'a) (y:'b): + card (fdom (m.[x<-y])) = if x \in m then card (fdom m) else card (fdom m) + 1. + proof. + rewrite fdom_set;case (x \in m)=> Hx. + + by rewrite fsetUC subset_fsetU_id 2:// => z; rewrite ?inE mem_fdom. + rewrite fcardUI_indep 1:fset0_eqP=>[z|]. + + by rewrite !inE;case (z=x)=>//; rewrite mem_fdom. + by rewrite fcard1. + qed. + + lemma size_set_le (m:('a,'b)fmap) (x:'a) (y:'b): card (fdom (m.[x<-y])) <= card (fdom m) + 1. + proof. rewrite size_set /#. qed. + + lemma size_rem (m:('a,'b)fmap) (x:'a): + card (fdom (rem m x)) = if x \in m then card (fdom m) - 1 else card (fdom m). + proof. + rewrite fdom_rem fcardD;case (x \in m)=> Hx. + + by rewrite subset_fsetI_id 2:fcard1// => z;rewrite !inE mem_fdom. + by rewrite (@eq_fcards0 (fdom m `&` fset1 x)) 2:// fset0_eqP=>z;rewrite !inE mem_fdom/#. + qed. + + lemma size_rem_le (m:('a,'b)fmap) (x:'a) : card (fdom (rem m x)) <= card (fdom m). + proof. by rewrite size_rem /#. qed. + + lemma size_ge0 (m:('a,'b)fmap) : 0 <= card (fdom m). + proof. rewrite fcard_ge0. qed. + + lemma size0 : card (fdom empty<:'a,'b>) = 0. + proof. by rewrite fdom0 fcards0. qed. + + local equiv RROset_inv_lt : RRO.set ~ RRO.set : + ={x, y, FRO.m} /\ inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2} ==> + ={res,FRO.m} /\ inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2}. + proof. + proc;auto=> &ml&mr[#]3!-> /= @/inv_lt [*]. + rewrite restr_set /=;smt w=(size_set_le size_rem_le). + qed. + + local lemma Pr_ext &m: + Pr[Gext.distinguish()@&m : G1.bext /\ ReSample.count <= max_size] <= + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + fel 8 ReSample.count (fun x=> (2*max_size)%r / (2^c)%r) + max_size G1.bext + [ReSample.f : + (card (fdom G1.m) <= max_size /\ card (fdom G1.mi) <= max_size /\ ReSample.count < max_size); + ReSample.f1 : + (card (fdom G1.m) < max_size /\ card (fdom G1.mi) < max_size /\ ReSample.count < max_size) + ]=> //; 2:by auto. + + rewrite /felsum Bigreal.sumr_const count_predT size_range. + apply ler_wpmul2r;1:by apply eps_ge0. + by rewrite le_fromint;smt ml=0 w=max_ge0. + + proc;rcondt 2;1:by auto. + wp; rnd (mem (image snd (fdom G1.m `|` fdom G1.mi ))); skip=> /> &hr h1 h2 h3 h4 h5. + rewrite (Mu_mem.mu_mem + (image snd (fdom G1.m{hr} `|` fdom G1.mi{hr})) + cdistr (1%r/(2^c)%r))//. + + by move=>x _;rewrite DCapacity.dunifin1E capacity_card. + rewrite ler_wpmul2r;1:by apply divr_ge0=>//;apply /c_ge0r. + rewrite imageU fcardU le_fromint. + move:(fcard_image_leq snd (fdom G1.m{hr}))(fcard_image_leq snd (fdom G1.mi{hr})). + smt(fcard_ge0). + + by move=>c1;proc;auto=> &hr [^H 2->]/#. + + by move=> b1 c1;proc;auto=> /#. + + proc;rcondt 2;1:by auto. + wp;rnd (mem (image snd (fdom G1.m `|` fdom G1.mi) `|` fset1 x));skip=> /> &hr ?????. + rewrite (Mu_mem.mu_mem (image snd (fdom G1.m{hr}`|`fdom G1.mi{hr}) `|` fset1 x{hr}) cdistr (1%r/(2^c)%r))//. + + by move=>x _;rewrite DCapacity.dunifin1E capacity_card. + rewrite ler_wpmul2r;1:by apply divr_ge0=>//;apply /c_ge0r. + rewrite imageU !fcardU le_fromint fcard1. + move:(fcard_image_leq snd (fdom G1.m{hr}))(fcard_image_leq snd (fdom G1.mi{hr})). + smt w=fcard_ge0. + + by move=>c1;proc;auto=> &hr [^H 2->]/#. + by move=> b1 c1;proc;auto=> /#. + qed. + + + local equiv EG2_Gext : Eager(G2(DRestr(D))).main2 ~ Gext.distinguish: + ={glob D} ==> + ReSample.count{2} <= max_size /\ + ((G1.bext{1} \/ inv_ext G1.m{1} G1.mi{1} FRO.m{1}) => G1.bext{2}). + proof. + proc;inline *;wp. + while (={l,FRO.m,G1.m,G1.mi} /\ card (fdom G1.m{2}) <= max_size /\ + card (fdom G1.mi{2}) <= max_size /\ + ReSample.count{2} + size l{2} <= max_size /\ + ((G1.bext{1} \/ + exists (x : state) (h : handle), + mem (fdom G1.m{1} `|` fdom G1.mi{1}) x /\ + FRO.m{1}.[h] = Some (x.`2, Unknown) /\ !mem l{1} h) => + G1.bext{2})). + + rcondt{2} 3. + + move=> &m;auto=> &m'[#] 6!-> /= + _ _;case (l{m'})=>//=; smt w=List.size_ge0. + auto=> &ml&mr[#]6!->;case(l{mr})=>[//|h1 l1/=Hle Hext c ?/=];split. + + smt w=(drop0 size_ge0). + rewrite drop0=>-[H|[x h][#]];1:by rewrite Hext // H. + rewrite get_setE;case (h=h1)=> [/=->Hin->_ | Hneq ???]. + + by right;apply (mem_image snd _ x). + by rewrite Hext 2://;right;exists x h;rewrite Hneq. + wp; call (_: ={F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext, C.c,C.queries} /\ + inv_le G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2}). + + proc;sp;if=>//=;swap -1. + call (_: ={x,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext,C.c,C.queries} /\ + inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2} ==> + ={res,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext,C.c,C.queries} /\ + inv_le G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2});last by auto=>/#. + proc;if=>//;last by auto=>/#. + seq 8 9 : (={x, y, F.RO.m, FRO.m, G1.paths, G1.mh, G1.mhi, G1.m, G1.mi, G1.chandle, + G1.bext, C.c, C.queries} /\ + inv_le G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2});2:by auto. + seq 2 3 : + (={y,x,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext, C.c,C.queries} /\ + inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2}). + + by if=>//;auto;call (_: ={F.RO.m});auto. + seq 5 5 : + (={t,y,x,hx2,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext, C.c,C.queries} /\ + inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2} /\ + (t => in_dom_with FRO.m (oget G1.mh.[(x.`1, hx2)]).`2 Unknown){1}). + + inline RRO.queried; wp;call (_: ={FRO.m});1:by sim. + inline RRO.allKnown;sp 1 1;if=>//. + by wp;call RROset_inv_lt;auto. + if=>//;wp. + + inline *;rcondt{1} 4;1:by auto=>/#. + rcondt{2} 5;1:by auto;smt w=(size_ge0). + rcondt{2} 10. by auto;progress;rewrite mem_set. + wp;rnd{2};auto=> /= ??[#]!-> @/inv_lt @/inv_le [#] mlt milt clt cle Hin 4? /=. + move => ? _;rewrite /bad_ext !get_setE /= set_set_eqE //=. + rewrite !(imageU,inE) restr_set /= size_rem dom_restr Hin //=; smt w=size_set_le. + by call RROset_inv_lt;auto;smt w=size_set_le. + + + proc;sp;if=> //;swap -1. + call (_: ={x,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext,C.c} /\ + inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2} ==> + ={res,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext,C.c} /\ + inv_le G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2});last by auto=> /#. + proc;if=>//;last by auto=>/#. + seq 8 8 : + (={t,y,x,hx2,F.RO.m,FRO.m,G1.paths,G1.mh,G1.mhi,G1.m,G1.mi,G1.chandle,G1.bext, C.c} /\ + inv_lt G1.m{2} G1.mi{2} C.c{1} FRO.m{2} ReSample.count{2} /\ + (t => in_dom_with FRO.m (oget G1.mhi.[(x.`1, hx2)]).`2 Unknown){1}). + + inline RRO.queried; auto;call (_: ={FRO.m});1:by sim. + inline RRO.allKnown;sp 1 1;if=>//. + by wp;call RROset_inv_lt;auto. + if=>//;wp. + + inline *;rcondt{1} 4;1:by auto=>/#. + rcondt{2} 5;1:by auto;smt w=(size_ge0). + rcondt{2} 10. by auto;progress;rewrite mem_set. + wp;rnd{2};auto=> /= ??[#]!-> @/inv_lt @/inv_le [#] mlt milt clt cle Hin 5? _. + rewrite /bad_ext !get_setE /= set_set_eqE //=. + rewrite !(imageU,inE) restr_set /= size_rem dom_restr Hin //=; smt w=size_set_le. + by call RROset_inv_lt;auto;smt w=size_set_le. + + + proc;sp 1 1;if;auto;if;auto=>//. + inline G2(DRestr(D), RRO).M.f Gext.M.f. + sp 6 6;elim *=> c0L c0R. + wp;call (_: ={F.RO.m});1:by sim. + conseq(:_==> ={i,p,G1.mh,sa,h,FRO.m,F.RO.m,G1.mh,G1.mhi,G1.chandle,counter} /\ + 0 <= i{1} <= size p{1} /\ + 0 <= counter{1} <= size p{1} - + prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) /\ + c0R + size p{1} - + prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) <= max_size /\ + inv_le G1.m{2} G1.mi{2} (c0R + counter){2} FRO.m{2} ReSample.count{2}). + + by auto=> />; smt(List.size_ge0). + while (={bs,i,p,G1.mh,sa,h,FRO.m,F.RO.m,G1.mh,G1.mhi,G1.chandle,counter,C.queries} /\ + bs{1} = p{1} /\ 0 <= i{1} <= size p{1} /\ + 0 <= counter{1} <= size p{1} - + prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) /\ + c0R + size p{1} - + prefix bs{1} (get_max_prefix bs{1} (elems (fdom C.queries{1}))) <= max_size /\ + inv_le G1.m{2} G1.mi{2} (c0R + counter){2} FRO.m{2} ReSample.count{2}); + last first. + + by auto=> />;smt(List.size_ge0 prefix_sizel). + if=> //;1:by auto=>/#. + if=> //;2:by auto=>/#. + auto;call (_: ={F.RO.m});1:by sim. + inline *;auto=> &ml &mr [#]!->@/inv_le Hi0[#] _ H_c_0 H_c_max H1 [#]H_size_m H_size_mi H_count H2 H3/=. + rewrite H3/==>H_nin_dom H_counter_prefix c;rewrite DCapacity.dunifin_fu/=. + case(G1.chandle{mr} \in FRO.m{mr})=>//=[/#|]H_handle_in_dom. + progress;..-3,-1: rewrite/#; by rewrite restr_set_eq size_set/=/#. + + auto;progress[delta];rewrite ?(size0,restr0,restr_set,rem0,max_ge0,-sizeE,-cardE) //=. + + smt(size_rem_le size0). + + smt(). + + smt(). + + smt(). + + by elim H7=>// [[x h] [#]];rewrite -memE mem_fdom dom_restr /in_dom_with domE=> _ ->. + by apply H10. + qed. + + declare axiom D_ll: + forall (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}), + islossless P.f => islossless P.fi => islossless F.f => islossless D(F, P).distinguish. + + lemma Real_G2 &m: + Pr[GReal(D).main() @ &m: res /\ C.c <= max_size] <= + Pr[Eager(G2(DRestr(D))).main2() @ &m: res] + + (max_size ^ 2 - max_size)%r / 2%r * mu dstate (pred1 witness) + + max_size%r * ((2*max_size)%r / (2^c)%r) + + max_size%r * ((2*max_size)%r / (2^c)%r). + proof. + apply (ler_trans _ _ _ (Real_G1 D D_ll &m)). + do !apply ler_add => //. + + have ->: Pr[G1(DRestr(D)).main() @ &m : res] = Pr[Eager(G2(DRestr(D))).main1() @ &m : res]. + + by byequiv (G1_G2 D). + by apply lerr_eq;byequiv (Eager_1_2 (G2(DRestr(D)))). + + by apply (Pr_G1col D D_ll &m). + apply (ler_trans Pr[Eager(G2(DRestr(D))).main1()@&m: G1.bext \/ inv_ext G1.m G1.mi FRO.m]). + + by byequiv (G1_G2 D)=>//#. + apply (ler_trans Pr[Eager(G2(DRestr(D))).main2()@&m : G1.bext \/ inv_ext G1.m G1.mi FRO.m]). + + by apply lerr_eq;byequiv (Eager_1_2 (G2(DRestr(D)))). + apply (ler_trans _ _ _ _ (Pr_ext &m)). + byequiv EG2_Gext=>//#. + qed. + +end section EXT. + + + diff --git a/sha3/proof/smart_counter/Handle.eca b/sha3/proof/smart_counter/Handle.eca new file mode 100644 index 0000000..4ba8b17 --- /dev/null +++ b/sha3/proof/smart_counter/Handle.eca @@ -0,0 +1,2718 @@ +pragma -oldip. pragma +implicits. +require import Core Int Real StdOrder Ring. +require import List FSet SmtMap Common SLCommon. +require import DProd Dexcepted. +require import PROM. +(*...*) import Capacity IntID IntOrder DCapacity. + +require (*--*) ConcreteF. + +clone export PROM.FullRO as ROhandle with + type in_t <- handle, + type out_t <- capacity, + op dout _ <- cdistr, + type d_in_t <- unit, + type d_out_t <- bool. +export ROhandle.FullEager. + +clone export ConcreteF as ConcreteF1. + +module G1(D:DISTINGUISHER) = { + var m, mi : smap + var mh, mhi : hsmap + var chandle : int + var paths : (capacity, block list * block) fmap + var bext, bcol : bool + + module M = { + + proc f(p : block list): block = { + var sa, sa', sc; + var h, i, counter <- 0; + sa <- b0; + sc <- c0; + while (i < size p ) { + if ((sa +^ nth witness p i, h) \in mh) { + (sa, h) <- oget mh.[(sa +^ nth witness p i, h)]; + } else { + if (counter < size p - prefix p (get_max_prefix p (elems (fdom C.queries)))) { + sc <$ cdistr; + bcol <- bcol \/ hinv FRO.m sc <> None; + sa' <@ F.RO.get(take (i+1) p); + sa <- sa +^ nth witness p i; + mh.[(sa,h)] <- (sa', chandle); + mhi.[(sa',chandle)] <- (sa, h); + (sa,h) <- (sa',chandle); + FRO.m.[chandle] <- (sc,Unknown); + chandle <- chandle + 1; + counter <- counter + 1; + } + } + i <- i + 1; + } + sa <@ F.RO.get(p); + return sa; + } + } + + module S = { + + proc f(x : state): state = { + var p, v, y, y1, y2, hy2, hx2; + + if (x \notin m) { + if (x.`2 \in paths) { + (p,v) <- oget paths.[x.`2]; + y1 <@ F.RO.get (rcons p (v +^ x.`1)); + y2 <$ cdistr; + } else { + y1 <$ bdistr; + y2 <$ cdistr; + } + y <- (y1, y2); + bext <- bext \/ rng FRO.m (x.`2, Unknown); + if (!(rng FRO.m (x.`2, Known))) { + FRO.m.[chandle] <- (x.`2, Known); + chandle <- chandle + 1; + } + hx2 <- oget (hinvK FRO.m x.`2); + if ((x.`1, hx2) \in mh /\ in_dom_with FRO.m (oget mh.[(x.`1,hx2)]).`2 Unknown) { + hy2 <- (oget mh.[(x.`1, hx2)]).`2; + y <- (y.`1, (oget FRO.m.[hy2]).`1); + FRO.m.[hy2] <- (y.`2, Known); + m.[x] <- y; + mi.[y] <- x; + } else { + bcol <- bcol \/ hinv FRO.m y.`2 <> None; + hy2 <- chandle; + chandle <- chandle + 1; + FRO.m.[hy2] <- (y.`2, Known); + m.[x] <- y; + mh.[(x.`1, hx2)] <- (y.`1, hy2); + mi.[y] <- x; + mhi.[(y.`1, hy2)] <- (x.`1, hx2); + } + if (x.`2 \in paths) { + (p,v) <- oget paths.[x.`2]; + paths.[y.`2] <- (rcons p (v +^ x.`1), y.`1); + } + } else { + y <- oget m.[x]; + } + return y; + } + + proc fi(x : state): state = { + var y, y1, y2, hx2, hy2; + + if (x \notin mi) { + bext <- bext \/ rng FRO.m (x.`2, Unknown); + if (!(rng FRO.m (x.`2, Known))) { + FRO.m.[chandle] <- (x.`2, Known); + chandle <- chandle + 1; + } + hx2 <- oget (hinvK FRO.m x.`2); + y1 <$ bdistr; + y2 <$ cdistr; + y <- (y1,y2); + if ((x.`1,hx2) \in mhi /\ + in_dom_with FRO.m (oget mhi.[(x.`1,hx2)]).`2 Unknown) { + (y1,hy2) <- oget mhi.[(x.`1, hx2)]; + y <- (y.`1, (oget FRO.m.[hy2]).`1); + FRO.m.[hy2] <- (y.`2, Known); + mi.[x] <- y; + m.[y] <- x; + } else { + bcol <- bcol \/ hinv FRO.m y.`2 <> None; + hy2 <- chandle; + chandle <- chandle + 1; + FRO.m.[hy2] <- (y.`2, Known); + mi.[x] <- y; + mhi.[(x.`1, hx2)] <- (y.`1, hy2); + m.[y] <- x; + mh.[(y.`1, hy2)] <- (x.`1, hx2); + } + } else { + y <- oget mi.[x]; + } + return y; + } + + } + + proc main(): bool = { + var b; + + F.RO.m <- empty; + m <- empty; + mi <- empty; + mh <- empty; + mhi <- empty; + bext <- false; + bcol <- false; + C.queries<- empty.[[] <- b0]; + + (* the empty path is initially known by the adversary to lead to capacity 0^c *) + FRO.m <- empty.[0 <- (c0, Known)]; + paths <- empty.[c0 <- ([<:block>],b0)]; + chandle <- 1; + b <@ D(M,S).distinguish(); + return b; + } +}. + +(* -------------------------------------------------------------------------- *) +(** The state of CF contains + - the map PF.m that represents the primitive's map. + - the map Redo.prefixes that contains all the prefixes computations of the + sponge construction. + The state of G1 contains: + - the map hs that associates handles to flagged capacities; + - the map G1.m that represents the *public* view of map PF.m; + - the map G1.mh that represents PF.m with handle-based indirection; + - the map ro that represents the functionality; + - the map pi that returns *the* known path to a capacity if it exists. + The following invariants encode these facts, and some auxiliary + knowledge that can most likely be deduced but is useful in the proof. **) + +(** RELATIONAL: Map, Handle-Map and Handles are compatible **) +inductive m_mh (hs : handles) (m : smap) (mh : hsmap) = + | INV_m_mh of (forall xa xc ya yc, + m.[(xa,xc)] = Some (ya,yc) => + exists hx fx hy fy, + hs.[hx] = Some (xc,fx) + /\ hs.[hy] = Some (yc,fy) + /\ mh.[(xa,hx)] = Some (ya,hy)) + & (forall xa hx ya hy, + mh.[(xa,hx)] = Some (ya,hy) => + exists xc fx yc fy, + hs.[hx] = Some (xc,fx) + /\ hs.[hy] = Some (yc,fy) + /\ m.[(xa,xc)] = Some (ya,yc)). + + +(* WELL-FORMEDNESS<1 >: Map and Prefixes are compatible *) +inductive m_p (m : smap) (p : (block list, state) fmap) + (q : (block list, block) fmap) = + | INV_m_p of (p.[[]] = Some (b0,c0)) + & (q.[[]] = Some b0) + & (forall (l : block list), + l \in p => + (forall i, 0 <= i < size l => + exists sa sc, p.[take i l] = Some (sa, sc) /\ + m.[(sa +^ nth witness l i, sc)] = p.[take (i+1) l])) + & (forall (l : block list), + l \in q => exists c, p.[l] = Some (oget q.[l], c)) + & (forall (l : block list), + l \in p => exists (l2 : block list), l ++ l2 \in q). + +(** RELATIONAL : Prefixes and RO are compatible. **) +inductive ro_p (ro : (block list, block) fmap) (p : (block list, state) fmap) = + | INV_ro_p of (ro = map (+ (fun (a:state)=> a.`1)) p). + +(* WELL-FORMEDNESS<2>: Handles, Map, Handle-Map and RO are compatible *) +inductive mh_spec (hs : handles) (Gm : smap) (mh : hsmap) (ro : (block list,block) fmap) = + | INV_mh of (forall xa hx ya hy, + mh.[(xa,hx)] = Some (ya,hy) => + exists xc fx yc fy, + hs.[hx] = Some (xc,fx) + /\ hs.[hy] = Some (yc,fy) + /\ if fy = Known + then Gm.[(xa,xc)] = Some (ya,yc) + /\ fx = Known + else exists p v, + ro.[rcons p (v +^ xa)] = Some ya + /\ build_hpath mh p = Some (v,hx)) + & (forall p bn b, + ro.[rcons p bn] = Some b <=> + exists v hx hy, + build_hpath mh p = Some (v,hx) + /\ mh.[(v +^ bn,hx)] = Some (b,hy)) + & (forall p v p' v' hx, + build_hpath mh p = Some (v,hx) + => build_hpath mh p' = Some (v',hx) + => p = p' /\ v = v'). + +(* WELL-FORMEDNESS<2>: Handles, Handle-Map and Paths are compatible *) +inductive pi_spec (hs : handles) (mh : hsmap) (pi : (capacity,block list * block) fmap) = + | INV_pi of (forall c p v, + pi.[c] = Some (p,v) <=> + exists h, + build_hpath mh p = Some(v,h) + /\ hs.[h] = Some (c,Known)). + +(* WELL-FORMEDNESS<2>: Handles are well-formed *) +inductive hs_spec hs ch = + | INV_hs of (huniq hs) + & (hs.[0] = Some (c0,Known)) + & (forall cf h, hs.[h] = Some cf => h < ch). + +(* Useless stuff *) +inductive inv_spec (m:('a,'b) fmap) mi = + | INV_inv of (forall x y, m.[x] = Some y <=> mi.[y] = Some x). + +(* Invariant: maybe we should split relational and non-relational parts? *) +inductive INV_CF_G1 (hs : handles) ch (Pm Pmi Gm Gmi : smap) + (mh mhi : hsmap) (ro : (block list,block) fmap) pi + (p : (block list, state) fmap) + (q : (block list, block) fmap) = + | HCF_G1 of (hs_spec hs ch) + & (inv_spec Gm Gmi) + & (inv_spec mh mhi) + & (m_mh hs Pm mh) + & (m_mh hs Pmi mhi) + & (incl Gm Pm) + & (incl Gmi Pmi) + & (mh_spec hs Gm mh ro) + & (pi_spec hs mh pi) + (* & (all_prefixes_fset (dom ro)) *) + & (m_p Pm p q). + +(** Structural Projections **) +lemma m_mh_of_INV (ch : handle) + (mi1 m2 mi2 : smap) (mhi2 : hsmap) + (ro : (block list, block) fmap) + (pi : (capacity, block list * block) fmap) + hs m1 mh2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + m_mh hs m1 mh2. +proof. by case. qed. + +lemma mi_mhi_of_INV (ch : handle) + (m1 m2 mi2 : smap) (mh2 : hsmap) + (ro : (block list, block) fmap) + (pi : (capacity, block list * block) fmap) + hs mi1 mhi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + m_mh hs mi1 mhi2. +proof. by case. qed. + +lemma incl_of_INV (hs : handles) (ch : handle) + (mi1 mi2 : smap) (mh2 mhi2: hsmap) + (ro : (block list, block) fmap) + (pi : (capacity, block list * block) fmap) + m1 m2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + incl m2 m1. +proof. by case. qed. + +lemma incli_of_INV (hs : handles) (ch : handle) + (m1 m2 : smap) (mh2 mhi2: hsmap) + (ro : (block list, block) fmap) + (pi : (capacity, block list * block) fmap) + mi1 mi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + incl mi2 mi1. +proof. by case. qed. + +lemma mh_of_INV (ch : handle) + (m1 mi1 mi2 : smap) (mhi2 : hsmap) + (pi : (capacity, block list * block) fmap) + hs m2 mh2 ro p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + mh_spec hs m2 mh2 ro. +proof. by case. qed. + +lemma pi_of_INV (ch : handle) + (m1 m2 mi1 mi2: smap) (mhi2: hsmap) + (ro : (block list, block) fmap) + hs mh2 pi p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + pi_spec hs mh2 pi. +proof. by case. qed. + +lemma hs_of_INV (m1 m2 mi1 mi2 : smap) (mh2 mhi2 : hsmap) + (ro : (block list, block) fmap) + (pi : (capacity, block list * block) fmap) + hs ch p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + hs_spec hs ch. +proof. by case. qed. + +lemma inv_of_INV hs ch m1 mi1 m2 mi2 ro pi + mh2 mhi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q=> + inv_spec mh2 mhi2. +proof. by case. qed. + +lemma invG_of_INV hs ch m1 mi1 mh2 mhi2 ro pi m2 mi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + inv_spec m2 mi2. +proof. by case. qed. + +(* lemma all_prefixes_fset_of_INV hs ch m1 mi1 mh2 mhi2 ro pi m2 mi2 p q: *) +(* INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => *) +(* all_prefixes_fset (dom ro). *) +(* proof. by case. qed. *) + +lemma m_p_of_INV hs ch m1 mi1 mh2 mhi2 ro pi m2 mi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q => + m_p m1 p q. +proof. by case. qed. + +lemma all_prefixes_of_m_p m1 p q: + m_p m1 p q => all_prefixes p. +proof. +case=>h0 h0' h1 h2 _ l hl i. +case(l = [])=>//=l_notnil. +case(0 <= i)=>hi0;last first. ++ by rewrite take_le0 1:/# domE h0. +case(i < size l)=>hisize;last smt(take_oversize). +smt(domE). +qed. + +lemma all_prefixes_of_INV hs ch m1 mi1 mh2 mhi2 ro pi m2 mi2 p q: + INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p q=> + all_prefixes p. +proof. case=>? ? ? ? ? ? ? ? ? h ?;exact(all_prefixes_of_m_p _ _ h). qed. + +(* lemma ro_p_of_INV hs ch m1 mi1 mh2 mhi2 ro pi m2 mi2 p: *) +(* INV_CF_G1 hs ch m1 mi1 m2 mi2 mh2 mhi2 ro pi p => *) +(* ro_p ro p. *) +(* proof. by case. qed. *) + +(** Useful Lemmas **) +lemma ch_gt0 hs ch : hs_spec hs ch => 0 < ch. +proof. by case=> _ + Hlt -/Hlt. qed. + +lemma ch_neq0 hs ch : hs_spec hs ch => 0 <> ch. +proof. by move=> /ch_gt0/ltr_eqF. qed. + +lemma ch_notdomE_hs hs ch: hs_spec hs ch => hs.[ch] = None. +proof. +by move=> [] _ _ dom_hs; case: {-1}(hs.[ch]) (eq_refl hs.[ch])=> [//|cf/dom_hs]. +qed. + +lemma Sch_notdomE_hs hs ch: hs_spec hs ch => hs.[ch + 1] = None. +proof. +by move=> [] _ _ dom_hs; case: {-1}(hs.[ch + 1]) (eq_refl hs.[ch + 1])=> [//|cg/dom_hs/#]. +qed. + +lemma ch_notdomE2_mh hs m mh xa ch: + m_mh hs m mh + => hs_spec hs ch + => mh.[(xa,ch)] = None. +proof. +move=> [] Hm_mh Hmh_m [] _ _ dom_hs. +case: {-1}(mh.[(xa,ch)]) (eq_refl mh.[(xa,ch)])=> [//=|[ya hy] /Hmh_m]. +by move=> [xc0 fx0 yc fy] [#] /dom_hs. +qed. + +lemma Sch_notdomE2_mh hs m mh xa ch: + m_mh hs m mh + => hs_spec hs ch + => mh.[(xa,ch + 1)] = None. +proof. +move=> [] Hm_mh Hmh_m [] _ _ dom_hs. +case: {-1}(mh.[(xa,ch + 1)]) (eq_refl mh.[(xa,ch + 1)])=> [//=|[ya hy] /Hmh_m]. +by move=> [xc0 fx0 yc fy] [#] /dom_hs /#. +qed. + +lemma dom_hs_neq_ch hs ch hx xc fx: + hs_spec hs ch + => hs.[hx] = Some (xc,fx) + => hx <> ch. +proof. by move=> [] _ _ dom_hs /dom_hs /#. qed. + +lemma dom_hs_neq_Sch hs ch hx xc fx: + hs_spec hs ch + => hs.[hx] = Some(xc,fx) + => hx <> ch + 1. +proof. by move=> [] _ _ dom_hs /dom_hs /#. qed. + +lemma notin_m_notin_mh hs m mh xa xc hx fx: + m_mh hs m mh + => m.[(xa,xc)] = None + => hs.[hx] = Some (xc,fx) + => mh.[(xa,hx)] = None. +proof. +move=> [] _ Hmh_m m_xaxc hs_hx; case: {-1}(mh.[(xa,hx)]) (eq_refl mh.[(xa,hx)])=> [//|]. +by move=> [ya hy] /Hmh_m [xc0 fx0 yc0 fy0] [#]; rewrite hs_hx=> [#] <*>; rewrite m_xaxc. +qed. + +lemma notin_m_notin_Gm (m Gm : ('a,'b) fmap) x: + incl Gm m + => m.[x] = None + => Gm.[x] = None. +proof. by move=> Gm_leq_m; apply/contraLR=> ^ /Gm_leq_m ->. qed. + +lemma notin_hs_notdomE2_mh hs m mh xa hx: + m_mh hs m mh + => hs.[hx] = None + => mh.[(xa,hx)] = None. +proof. +move=> [] _ Hmh_m hs_hx; case: {-1}(mh.[(xa,hx)]) (eq_refl mh.[(xa,hx)])=> [//|]. +by move=> [ya hy] /Hmh_m [xc fx yc fy] [#]; rewrite hs_hx. +qed. + +(** Preservation of m_mh **) +lemma m_mh_addh hs ch m mh xc fx: + hs_spec hs ch + => m_mh hs m mh + => m_mh hs.[ch <- (xc, fx)] m mh. +proof. +move=> ^Hhs [] Hhuniq hs_0 dom_hs [] Hm_mh Hmh_m; split. ++ move=> xa0 xc0 ya yc /Hm_mh [hx0 fx0 hy fy] [#] hs_hx0 hs_hy mh_xaxc0. + exists hx0 fx0 hy fy; rewrite !get_setE mh_xaxc0 hs_hx0 hs_hy /=. + move: hs_hx0=> /dom_hs/ltr_eqF -> /=. + by move: hs_hy=> /dom_hs/ltr_eqF -> /=. +move=> xa hx ya hy /Hmh_m [xc0 fx0 yc fy] [#] hs_hx hs_hy m_xaxc0. +exists xc0 fx0 yc fy; rewrite !get_setE m_xaxc0 hs_hx hs_hy. +move: hs_hx=> /dom_hs/ltr_eqF -> /=. +by move: hs_hy=> /dom_hs/ltr_eqF -> /=. +qed. + +lemma m_mh_updh fy0 hs m mh yc hy fy: + m_mh hs m mh + => hs.[hy] = Some (yc,fy0) + => m_mh hs.[hy <- (yc,fy)] m mh. +proof. +move=> Im_mh hs_hy; split. ++ move=> xa' xc' ya' yc'; have [] H _ /H {H}:= Im_mh. + move=> [hx' fx' hy' fy'] [#] hs_hx' hs_hy' mh_xahx'. + case: (hx' = hy); case: (hy' = hy)=> //= <*> => [|Hhy'|Hhx'|Hhx' Hhy']. + + by exists hy fy hy fy; rewrite !get_setE /= /#. + + by exists hy fy hy' fy'; rewrite !get_setE Hhy' /#. + + by exists hx' fx' hy fy; rewrite !get_setE Hhx' /#. + by exists hx' fx' hy' fy'; rewrite !get_setE Hhx' Hhy'. +move=> xa' hx' ya' hy'; have [] _ H /H {H}:= Im_mh. +move=> [xc' fx' yc' fy'] [#] hs_hx' hs_hy' m_xaxc'. +case: (hx' = hy); case: (hy' = hy)=> //= <*> => [|Hhy'|Hhx'|Hhx' Hhy']. ++ by exists yc fy yc fy; rewrite !get_setE /= /#. ++ by exists yc fy yc' fy'; rewrite !get_setE Hhy' /#. ++ by exists xc' fx' yc fy; rewrite !get_setE Hhx' /#. +by exists xc' fx' yc' fy'; rewrite !get_setE Hhx' Hhy'. +qed. + +lemma m_mh_addh_addm hs Pm mh hx xa xc hy ya yc f f': + m_mh hs Pm mh => + huniq hs => + hs.[hx] = Some (xc, f) => + hs.[hy] = None => + m_mh hs.[hy <- (yc,f')] Pm.[(xa,xc) <- (ya,yc)] mh.[(xa,hx) <- (ya,hy)]. +proof. +move=> [] Hm_mh Hmh_m Hhuniq hs_hx hs_hy. +split=> [xa0 xc0 ya0 yc0|xa0 hx0 ya0 hy0]; rewrite get_setE. ++ case: ((xa0,xc0) = (xa,xc))=> [[#] <<*> [#] <<*>|] /=. + + by exists hx f hy f'; rewrite !get_setE /= /#. + move=> xaxc0_neq_xaxc /Hm_mh [hx0 fx0 hy0 fy0] [#] hs_hx0 hs_hy0 mh_xahx0. + by exists hx0 fx0 hy0 fy0; rewrite !get_setE /#. +case: ((xa0,hx0) = (xa,hx))=> [[#] <*>> [#] <<*>|] /=. ++ by exists xc f yc f'; rewrite !get_setE /= /#. +move=> /= /negb_and xahx0_neq_xahx /Hmh_m [xc0 fx0 yc0 fy0] [#] hs_hx0 hs_hy0 Pm_xaxc0. +exists xc0 fx0 yc0 fy0; rewrite !get_setE; do !split=> [/#|/#|/=]. +move: xahx0_neq_xahx; case: (xa0 = xa)=> [/= <*>>|//=]; case: (xc0 = xc)=> [<*>>|//=]. +by move: hs_hx=> /(Hhuniq _ _ _ _ hs_hx0). +qed. + +lemma mi_mhi_addh_addmi (hs : handles) mi mhi hx xa xc hy ya yc fx fy: + m_mh hs mi mhi => + (forall f h, hs.[h] <> Some (yc,f)) => + hs.[hx] = Some (xc,fx) => + hs.[hy] = None => + m_mh hs.[hy <- (yc,fy)] mi.[(ya,yc) <- (xa,xc)] mhi.[(ya,hy) <- (xa,hx)]. +proof. +move=> [] Hm_mh Hmh_m yc_notin_rng1_hs hs_hx hs_hy; split. ++ move=> ya0 yc0 xa0 xc0; rewrite get_setE; case: ((ya0,yc0) = (ya,yc))=> [[#] <*>> [#] <*>>|]. + + by exists hy fy hx fx; rewrite !get_setE /= /#. + move=> yayc0_neq_yayc /Hm_mh [hy0 fy0 hx0 fx0] [#] hs_hy0 hs_hx0 mhi_yayc0. + by exists hy0 fy0 hx0 fx0; rewrite !get_setE /#. +move=> ya0 hy0 xa0 hx0; rewrite get_setE; case: ((ya0,hy0) = (ya,hy))=> [[#] <*>> [#] <<*>|]. ++ by exists yc fy xc fx; rewrite !get_setE //= /#. +move=>yahy0_neq_yahy /Hmh_m [yc0 fy0 xc0 fx0] [#] hs_hy0 hs_hx0 mi_yayc0. +exists yc0 fy0 xc0 fx0; rewrite !get_setE; do !split=> [/#|/#|]. +move: yahy0_neq_yahy; case: (ya0 = ya)=> [<<*> //=|/#]; case: (yc0 = yc)=> [<*>> /=|//=]. +by move: hs_hy0; rewrite yc_notin_rng1_hs. +qed. + +(** Inversion **) +lemma inv_mh_inv_Pm hs Pm Pmi mh mhi: + m_mh hs Pm mh + => m_mh hs Pmi mhi + => inv_spec mh mhi + => inv_spec Pm Pmi. +proof. +move=> Hm_mh Hmi_mhi [] Hinv; split=>- [xa xc] [ya yc]; split. ++ have [] H _ /H {H} [hx fx hy fy] [#] hs_hx hs_hy /Hinv := Hm_mh. + have [] _ H /H {H} [? ? ? ?] [#] := Hmi_mhi. + by rewrite hs_hx hs_hy=> /= [#] <<*> [#] <<*>. +have [] H _ /H {H} [hy fy hx fx] [#] hs_hy hs_hx /Hinv := Hmi_mhi. +have [] _ H /H {H} [? ? ? ?] [#] := Hm_mh. +by rewrite hs_hx hs_hy=> /= [#] <<*> [#] <<*>. +qed. + +lemma inv_incl_none Pm Pmi Gm (x : 'a) Gmi (y : 'b): + inv_spec Pm Pmi + => inv_spec Gm Gmi + => incl Gm Pm + => incl Gmi Pmi + => Pm.[x] = Some y + => (Gm.[x] = None <=> Gmi.[y] = None). +proof. +move=> [] invP [] invG Gm_leq_Pm Gmi_leq_Pmi ^P_x; rewrite invP=> Pi_y. +split=> [G_x | Gi_y]. ++ case: {-1}(Gmi.[y]) (eq_refl Gmi.[y])=> [//|x']. + move=> ^Gmi_y; rewrite -Gmi_leq_Pmi 1:Gmi_y// Pi_y /= -negP=> <<*>. + by move: Gmi_y; rewrite -invG G_x. +case: {-1}(Gm.[x]) (eq_refl Gm.[x])=> [//|y']. +move=> ^Gm_y; rewrite -Gm_leq_Pm 1:Gm_y// P_x /= -negP=> <<*>. +by move: Gm_y; rewrite invG Gi_y. +qed. + +(** Preservation of hs_spec **) +lemma huniq_addh hs h c f: + huniq hs + => (forall f' h', hs.[h'] <> Some (c,f')) + => huniq hs.[h <- (c,f)]. +proof. +move=> Hhuniq c_notin_rng1_hs h1 h2 [c1 f1] [c2 f2]; rewrite !get_setE. +case: (h1 = h); case: (h2 = h)=> //= [Hh2 + [#]|+ Hh1 + [#]|_ _] - <*>. ++ by rewrite c_notin_rng1_hs. ++ by rewrite c_notin_rng1_hs. +move=> H1 H2. +by have/=:=Hhuniq _ _ _ _ H1 H2. +qed. + +lemma hs_addh hs ch xc fx: + hs_spec hs ch + => (forall f h, hs.[h] <> Some (xc,f)) + => hs_spec hs.[ch <- (xc,fx)] (ch + 1). +proof. +move=> ^Hhs [] Hhuniq hs_0 dom_hs xc_notin_rng1_hs; split. ++ move=> h1 h2 [c1 f1] [c2 f2]; rewrite !get_setE /=. + case: (h1 = ch); case: (h2 = ch)=> //= [+ + [#]|+ + + [#]|]=> <*>; + first 2 by rewrite xc_notin_rng1_hs. + by move=> _ _ hs_h1 /(Hhuniq _ _ _ _ hs_h1). ++ by rewrite get_setE (ch_neq0 _ Hhs). ++ move=> [c f] h; rewrite !get_setE; case: (h = ch)=> [<*> /#|_]. +by move=> /dom_hs /#. +qed. + +lemma hs_updh hs ch fx hx xc fx': + hs_spec hs ch + => 0 <> hx + => hs.[hx] = Some (xc,fx) + => hs_spec hs.[hx <- (xc,fx')] ch. +proof. +move=> ^Hhs [] Hhuniq hs_0 dom_hs hx_neq0 hs_hx; split. ++ move=> h1 h2 [c1 f1] [c2 f2]; rewrite !get_setE /= => />. + case : (h1 = hx) => />; case : (h2 = hx) => /> U1 U2. + + by have := (Hhuniq _ _ _ _ hs_hx U2 _). + + case (xc = c2) => />. + by have := (Hhuniq _ _ _ _ hs_hx U2 _) => // />. + + move=> H1 H2; by have := (Hhuniq _ _ _ _ H1 H2 _). ++ by rewrite get_setE hx_neq0. +move=> cf h; rewrite get_setE; case: (h = hx)=> [<*> _|_ /dom_hs //]. +by move: hs_hx=> /dom_hs. +qed. + +(** Preservation of mh_spec **) +lemma mh_addh hs ch Gm mh ro xc fx: + hs_spec hs ch + => mh_spec hs Gm mh ro + => mh_spec hs.[ch <- (xc,fx)] Gm mh ro. +proof. +move=> [] _ _ dom_hs [] Hmh ? ?; split=> //. +move=> xa hx ya hy /Hmh [xc0 fx0 yc0 fy0] [#] hs_hx hs_hy Hite. +exists xc0 fx0 yc0 fy0; rewrite !get_setE Hite hs_hx hs_hy /=. +rewrite ltr_eqF /=; 1:by apply/(dom_hs _ hs_hx). +by rewrite ltr_eqF /=; 1:by apply/(dom_hs _ hs_hy). +qed. + +(** Preservation of inv_spec **) +lemma inv_addm (m : ('a,'b) fmap) mi x y: + inv_spec m mi + => m.[x] = None + => mi.[y] = None + => inv_spec m.[x <- y] mi.[y <- x]. +proof. +move=> [] Hinv m_x mi_y; split=> x' y'; rewrite !get_setE; split. ++ case: (x' = x)=> /= [[#] <*> //=|_ /Hinv ^ + ->]. + by move: mi_y; case: (y' = y)=> [[#] <*> ->|]. +case: (y' = y)=> /= [[#] <*> //=|_ /Hinv ^ + ->]. +by move: m_x; case: (x' = x)=> [[#] <*> ->|]. +qed. + +(** Preservation of incl **) +lemma incl_addm (m m' : ('a,'b) fmap) x y: + incl m m' + => incl m.[x <- y] m'.[x <- y]. +proof. by move=> m_leq_m' x'; rewrite !get_setE; case: (x' = x)=> [|_ /m_leq_m']. qed. + +(** getflag: retrieve the flag of a capacity **) +op getflag (hs : handles) xc = + omap snd (obind ("_.[_]" hs) (hinv hs xc)). + +lemma getflagP_none hs xc: + (getflag hs xc = None <=> forall f h, hs.[h] <> Some (xc,f)). +proof. by rewrite /getflag; case: (hinvP hs xc)=> [->|] //= /#. qed. + +lemma getflagP_some hs xc f: + huniq hs + => (getflag hs xc = Some f <=> rng hs (xc,f)). +proof. +move=> huniq_hs; split. ++ rewrite /getflag; case: (hinvP hs xc)=> [-> //|]. + rewrite rngE; case: (hinv hs xc)=> //= h [f'] ^ hs_h -> @/snd /= ->>. + by exists h. +rewrite rngE=> -[h] hs_h. +move: (hinvP hs xc)=> [_ /(_ h f) //|]. +rewrite /getflag; case: (hinv hs xc)=> // h' _ [f']; rewrite oget_some. +move=> /(huniq_hs _ h _ (xc,f)) /(_ hs_h) /= ->>. +by rewrite hs_h. +qed. + +(** Stuff about paths **) +lemma build_hpath_prefix mh p b v h: + build_hpath mh (rcons p b) = Some (v,h) + <=> (exists v' h', build_hpath mh p = Some (v',h') /\ mh.[(v' +^ b,h')] = Some (v,h)). +proof. +rewrite build_hpathP; split=> [[|p' b' v' h' [#] + Hhpath Hmh]|[v' h'] [] Hhpath Hmh]. ++ smt(size_rcons size_ge0). ++ by move=> ^/rconsIs <<- /rconssI <<-; exists v' h'. +exact/(Extend _ _ _ _ _ Hhpath Hmh). +qed. + +lemma build_hpath_up mh xa hx ya hy p za hz: + build_hpath mh p = Some (za,hz) + => mh.[(xa,hx)] = None + => build_hpath mh.[(xa,hx) <- (ya,hy)] p = Some (za,hz). +proof. +move=> + mh_xahx; elim/last_ind: p za hz=> [za hz|p b ih za hz]. ++ by rewrite /build_hpath. +move=> /build_hpath_prefix [b' h'] [#] /ih Hpath Hmh. +apply/build_hpathP/(@Extend _ _ _ _ p b b' h' _ Hpath _)=> //. +by rewrite get_setE /#. +qed. + +lemma build_hpath_down mh xa hx ya hy p v h: + (forall p v, build_hpath mh p <> Some (v,hx)) + => build_hpath mh.[(xa,hx) <- (ya,hy)] p = Some (v,h) + => build_hpath mh p = Some (v,h). +proof. +move=> no_path_to_hx. +elim/last_ind: p v h=> [v h /build_hpathP [<*>|] //=|p b ih]. ++ smt(size_ge0 size_rcons). +move=> v h /build_hpathP [/#|p' b' + + ^/rconsIs <<- /rconssI <<-]. +move=> v' h' /ih; rewrite get_setE. +case: ((v' +^ b,h') = (xa,hx))=> [/#|_ Hpath Hextend]. +exact/build_hpathP/(Extend _ _ _ _ _ Hpath Hextend). +qed. + +lemma known_path_uniq hs mh pi xc hx p xa p' xa': + pi_spec hs mh pi + => hs.[hx] = Some (xc,Known) + => build_hpath mh p = Some (xa, hx) + => build_hpath mh p' = Some (xa',hx) + => p = p' /\ xa = xa'. +proof. +move=> [] Ipi hs_hy path_p path_p'. +have /iffRL /(_ _):= Ipi xc p xa; first by exists hx. +have /iffRL /(_ _):= Ipi xc p' xa'; first by exists hx. +by move=> ->. +qed. + +(* Useful? Not sure... *) +lemma path_split hs ch m mh xc hx p xa: + hs_spec hs ch + => m_mh hs m mh + => hs.[hx] = Some (xc,Unknown) + => build_hpath mh p = Some (xa,hx) + => exists pk ya yc hy b za zc hz pu, + p = (rcons pk b) ++ pu + /\ build_hpath mh pk = Some (ya,hy) + /\ hs.[hy] = Some (yc,Known) + /\ mh.[(ya +^ b,hy)] = Some (za,hz) + /\ hs.[hz] = Some (zc,Unknown). +proof. +move=> Ihs [] _ Imh_m. +elim/last_ind: p hx xa xc=> [hx xa xc + /build_hpathP [_ <*>|]|]. ++ by have [] _ -> _ [#]:= Ihs. ++ smt(size_ge0 size_rcons). +move=> p b ih hx xa xc hs_hx /build_hpath_prefix. +move=> [ya hy] [#] path_p_hy ^mh_yabh' /Imh_m [yc fy ? ?] [#] hs_hy. +rewrite hs_hx=> /= [#] <<*> _; case: fy hs_hy. ++ move=> /ih /(_ ya _) // [pk ya' yc' hy' b' za zc hz pu] [#] <*>. + move=> Hpath hs_hy' mh_tahy' hs_hz. + by exists pk ya' yc' hy' b' za zc hz (rcons pu b); rewrite rcons_cat. +by move=> hs_hy; exists p ya yc hy b xa xc hx []; rewrite cats0. +qed. + +(** Path-specific lemmas **) +lemma lemma1 hs ch Pm Pmi Gm Gmi mh mhi ro pi x1 x2 y1 y2 prefixes queries: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries + => x2 <> y2 + => Pm.[(x1,x2)] = None + => Gm.[(x1,x2)] = None + => (forall f h, hs.[h] <> Some (x2,f)) + => (forall f h, hs.[h] <> Some (y2,f)) + => INV_CF_G1 + hs.[ch <- (x2,Known)].[ch + 1 <- (y2,Known)] (ch + 2) + Pm.[(x1,x2) <- (y1,y2)] Pmi.[(y1,y2) <- (x1,x2)] + Gm.[(x1,x2) <- (y1,y2)] Gmi.[(y1,y2) <- (x1,x2)] + mh.[(x1,ch) <- (y1,ch + 1)] mhi.[(y1,ch + 1) <- (x1,ch)] + ro pi prefixes queries. +proof. +move=> HINV x2_neq_y2 Pm_x Gm_x x2_notrngE1_hs y2_notrngE1_hs; split. ++ rewrite (@addzA ch 1 1); apply/hs_addh. + + by move: HINV=> /hs_of_INV/hs_addh=> ->. + by move=> f h; rewrite get_setE; case: (h = ch)=> [/#|_]; exact/y2_notrngE1_hs. ++ apply/inv_addm=> //; 1:by case: HINV. + case: {-1}(Gmi.[(y1,y2)]) (eq_refl Gmi.[(y1,y2)])=> [//|[xa xc]]. + + have /incli_of_INV @/incl + ^h - <- := HINV; 1: by rewrite h. + have /mi_mhi_of_INV [] H _ /H {H} [hx fx hy fy] [#] := HINV. + by rewrite y2_notrngE1_hs. ++ apply/inv_addm; 1:by case: HINV. + + have ^ /m_mh_of_INV Hm_mh /hs_of_INV Hhs := HINV. + by apply/(ch_notdomE2_mh _ _ Hm_mh Hhs). + have ^ /mi_mhi_of_INV Hmi_mhi /hs_of_INV Hhs := HINV. + by apply/(Sch_notdomE2_mh _ _ Hmi_mhi Hhs). ++ apply/(@m_mh_addh_addm hs.[ch <- (x2,Known)] Pm mh ch x1 x2 (ch + 1) y1 y2 Known). + + by move: HINV=> ^/hs_of_INV Hhs /m_mh_of_INV; exact/(m_mh_addh Hhs). + + by move: HINV => /hs_of_INV /hs_addh /(_ x2 Known _) // []. + + by rewrite get_setE. + by rewrite get_setE gtr_eqF 1:/# /=; apply/Sch_notdomE_hs; case: HINV. ++ apply/(@mi_mhi_addh_addmi hs.[ch <- (x2,Known)] Pmi mhi ch x1 x2 (ch + 1) y1 y2 Known Known). + + by move: HINV=> ^/hs_of_INV Hhs /mi_mhi_of_INV; exact/(m_mh_addh Hhs). + + move=> f h; rewrite get_setE; case: (h = ch)=> [_ //=|_ //=]; first by rewrite x2_neq_y2. + by rewrite y2_notrngE1_hs. + + by rewrite get_setE. + by rewrite get_setE gtr_eqF 1:/# /=; apply/Sch_notdomE_hs; case: HINV. ++ by apply/incl_addm; case: HINV. ++ by apply/incl_addm; case: HINV. ++ split. + + move=> xa hx ya hy; rewrite get_setE; case: ((xa,hx) = (x1,ch))=> [|]. + + by move=> [#] <*> [#] <*>; exists x2 Known y2 Known; rewrite !get_setE /#. + move=> xahx_neq_x1ch; have ^ /hs_of_INV Hhs /mh_of_INV [] Hmh _ _ /Hmh {Hmh} := HINV. + move=> [xc fx yc fy] [#] hs_hx hs_hy Hite. + exists xc fx yc fy; do 2?split; first 2 by smt (dom_hs_neq_ch dom_hs_neq_Sch get_setE). + case: fy Hite hs_hy=> /= [[p v] [Hro Hpath] hs_hy|[#] Gm_xaxc <*> hs_hy] /=; last first. + + by rewrite get_setE; case: ((xa,xc) = (x1,x2))=> [/#|]. + exists p v; rewrite Hro /=; apply/build_hpath_up=> //. + have /m_mh_of_INV /notin_hs_notdomE2_mh H:= HINV. + exact/H/ch_notdomE_hs/Hhs. + + move=> p xa b; have /mh_of_INV [] _ -> _ := HINV. + apply/exists_iff=> v /=; apply/exists_iff=> hx /=; apply/exists_iff=> hy /=. + have mh_x1ch: mh.[(x1,ch)] = None. + + by apply/(notin_hs_notdomE2_mh hs Pm)/ch_notdomE_hs; case: HINV. + have ch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch). + + move=> a h a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [xc fx yc fy] [#] _; rewrite ch_notdomE_hs; case: HINV. + split=> -[#]. + + move=> Hpath mh_vxahx; rewrite get_setE; case: ((v +^ xa,hx) = (x1,ch))=> [/#|_]. + by rewrite mh_vxahx //=; apply/build_hpath_up=> //=; rewrite mh_x1ch. + have H /H {H}:= build_hpath_down mh x1 ch y1 (ch + 1) p v hx _. + + move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} := HINV. + by move=> p' b' v' h' <*>; rewrite ch_notrngE2_mh. + move=> ^ /build_hpathP + -> /=; rewrite get_setE. + by case=> [<*>|/#]; move: HINV=> /hs_of_INV [] _ + H - /H {H} /#. + move=> p v p' v' hx. + have: (forall p v, build_hpath mh p <> Some (v,ch)). + + move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} := HINV. + move=> p'0 b'0 v'0 h'0 <*> _; have /m_mh_of_INV [] _ H /H {H} := HINV. + by move=> [xc fx yc fy] [#] _; have /hs_of_INV [] _ _ H /H {H}:= HINV. + move=> ^ + /build_hpath_down H /H {H} - /build_hpath_down H + /H {H}. + by have /mh_of_INV [] _ _ /(_ p v p' v' hx) := HINV. +split=> c p v; have ^/hs_of_INV [] _ _ dom_hs /pi_of_INV [] -> := HINV. +apply/exists_iff=> h /=; split=> [#]. ++ move=> /(build_hpath_up mh x1 ch y1 (ch + 1) p v h) /(_ _). + + by apply/(notin_hs_notdomE2_mh hs Pm)/ch_notdomE_hs; case: HINV. + by move=> -> /= ^ /dom_hs; rewrite !get_setE /#. +have ch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch). ++ move=> a h' a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [xc fx yc fy] [#] _; rewrite ch_notdomE_hs; case: HINV. +have Sch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch + 1). ++ move=> a h' a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [xc fx yc fy] [#] _; rewrite Sch_notdomE_hs; case: HINV. +have H /H {H}:= build_hpath_down mh x1 ch y1 (ch + 1) p v h _. ++ move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} := HINV. + by move=> p' b' v' h' <*>; rewrite ch_notrngE2_mh. ++ move=> ^ /build_hpathP + -> /=; rewrite !get_setE. + by case=> [<*>|/#]; move: HINV=> /hs_of_INV [] _ + H - /H {H} /#. +(* + by apply(ro_p_of_INV _ _ _ _ _ _ _ _ _ HINV). *) +split=>[]. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ [] ->//. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ [] //. ++ move=>l hmem i hi. + have[]_ _ h2 h3:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have[]sa sc[]:=h2 l hmem i hi. + have h1:=all_prefixes_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + smt(domE get_setE). +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +qed. + + +lemma lemma1' hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries x1 x2 y1 y2: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries + => ! (y1,y2) \in Pm + => x2 <> y2 + => Pmi.[(x1,x2)] = None + => Gmi.[(x1,x2)] = None + => (forall f h, hs.[h] <> Some (x2,f)) + => (forall f h, hs.[h] <> Some (y2,f)) + => INV_CF_G1 + hs.[ch <- (x2,Known)].[ch + 1 <- (y2,Known)] (ch + 2) + Pm.[(y1,y2) <- (x1,x2)] Pmi.[(x1,x2) <- (y1,y2)] + Gm.[(y1,y2) <- (x1,x2)] Gmi.[(x1,x2) <- (y1,y2)] + mh.[(y1,ch + 1) <- (x1,ch)] mhi.[(x1,ch) <- (y1,ch + 1)] + ro pi prefixes queries. +proof. +move=> HINV hh x2_neq_y2 Pm_x Gm_x xc_notrngE1_hs yc_notrngE1_hs; split. ++ rewrite (@addzA ch 1 1); apply/hs_addh. + + by move: HINV=> /hs_of_INV/hs_addh=> ->. + by move=> f h; rewrite get_setE; case: (h = ch)=> [/#|_]; exact/yc_notrngE1_hs. ++ apply/inv_addm=> //; 1:by case: HINV. + case: {-1}(Gm.[(y1,y2)]) (eq_refl Gm.[(y1,y2)])=> [//|[xa xc]]. + + have /incl_of_INV + ^h - <- := HINV; 1: by rewrite h. + have /m_mh_of_INV [] H _ /H {H} [hx fx hy fy] [#] := HINV. + by rewrite yc_notrngE1_hs. ++ apply/inv_addm; 1:by case: HINV. + + have ^ /m_mh_of_INV Hm_mh /hs_of_INV Hhs := HINV. + by apply/(Sch_notdomE2_mh _ _ Hm_mh Hhs). + have ^ /mi_mhi_of_INV Hmi_mhi /hs_of_INV Hhs := HINV. + by apply/(ch_notdomE2_mh _ _ Hmi_mhi Hhs). ++ apply/(@mi_mhi_addh_addmi hs.[ch <- (x2,Known)] Pm mh ch x1 x2 (ch + 1) y1 y2 Known Known). + + by move: HINV=> ^/hs_of_INV Hhs /m_mh_of_INV; exact/(m_mh_addh Hhs). + + by move=> f h; rewrite get_setE; case: (h = ch)=> [<*> /#|]; rewrite yc_notrngE1_hs. + + by rewrite get_setE. + by rewrite get_setE gtr_eqF 1:/# /=; apply/Sch_notdomE_hs; case: HINV. ++ apply/(@m_mh_addh_addm hs.[ch <- (x2,Known)] Pmi mhi ch x1 x2 (ch + 1) y1 y2 Known). + + by move: HINV=> ^/hs_of_INV Hhs /mi_mhi_of_INV; exact/(m_mh_addh Hhs). + + by have /hs_of_INV /hs_addh /(_ x2 Known _) // []:= HINV. + + by rewrite get_setE. + by rewrite get_setE gtr_eqF 1:/# /=; apply/Sch_notdomE_hs; case: HINV. ++ by apply/incl_addm; case: HINV. ++ by apply/incl_addm; case: HINV. ++ split. + + move=> ya hy xa hx; rewrite get_setE; case: ((ya,hy) = (y1,ch + 1))=> [|]. + + by move=> [#] <*> [#] <*>; exists y2 Known x2 Known; rewrite !get_setE /#. + move=> yahy_neq_y1Sch; have ^ /hs_of_INV Hhs /mh_of_INV [] Hmh _ _ /Hmh {Hmh} := HINV. + move=> [yc fy xc fx] [#] hs_hy hs_hx Hite. + exists yc fy xc fx; do 2?split; first 2 by smt (dom_hs_neq_ch dom_hs_neq_Sch get_setE). + case: fx Hite hs_hx=> /= [[p v] [Hro Hpath] hs_hx|[#] Gm_yayc <*> hs_hx] /=; last first. + + by rewrite get_setE; case: ((ya,yc) = (y1,y2))=> [/#|]. + exists p v; rewrite Hro /=; apply/build_hpath_up=> //. + have /m_mh_of_INV /notin_hs_notdomE2_mh H:= HINV. + exact/H/Sch_notdomE_hs/Hhs. + + move=> p ya b; have /mh_of_INV [] _ -> _ := HINV. + apply/exists_iff=> v /=; apply/exists_iff=> hx /=; apply/exists_iff=> hy /=. + have mh_y1Sch: mh.[(y1,ch + 1)] = None. + + by apply/(notin_hs_notdomE2_mh hs Pm)/Sch_notdomE_hs; case: HINV. + have Sch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch + 1). + + move=> a h a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [yc fy xc fx] [#] _; rewrite Sch_notdomE_hs; case: HINV. + split=> -[#]. + + move=> Hpath mh_vxahx; rewrite get_setE; case: ((v +^ ya,hx) = (y1,ch + 1))=> [/#|_]. + by rewrite mh_vxahx //=; apply/build_hpath_up=> //=; rewrite mh_y1ch. + have H /H {H}:= build_hpath_down mh y1 (ch + 1) x1 ch p v hx _. + + move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} /# := HINV. + by move=> p' b' v' h' <*>; rewrite Sch_notrngE2_mh. + move=> ^ /build_hpathP + -> /=; rewrite get_setE. + by case=> [<*>|/#]; move: HINV=> /hs_of_INV [] _ + H - /H {H} /#. + move=> p v p' v' hx. + have: (forall p v, build_hpath mh p <> Some (v,ch + 1)). + + move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} /# := HINV. + move=> p'0 b'0 v'0 h'0 <*> _; have /m_mh_of_INV [] _ H /H {H} := HINV. + by move=> [xc fx yc fy] [#] _; have /hs_of_INV [] _ _ H /H {H} /#:= HINV. + move=> ^ + /build_hpath_down H /H {H} - /build_hpath_down H + /H {H}. + by have /mh_of_INV [] _ _ /(_ p v p' v' hx) := HINV. +split=> c p v; have ^/hs_of_INV [] _ _ dom_hs /pi_of_INV [] -> := HINV. +apply/exists_iff=> h /=; split=> [#]. ++ move=> /(build_hpath_up mh y1 (ch + 1) x1 ch p v h) /(_ _). + + by apply/(notin_hs_notdomE2_mh hs Pm)/Sch_notdomE_hs; case: HINV. + by move=> -> /= ^ /dom_hs; rewrite !get_setE /#. +have ch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch). ++ move=> a h' a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [xc fx yc fy] [#] _; rewrite ch_notdomE_hs; case: HINV. +have Sch_notrngE2_mh: forall a h a', mh.[(a,h)] <> Some (a',ch + 1). ++ move=> a h' a'; rewrite -negP; have /m_mh_of_INV [] _ Hmh_m /Hmh_m {Hmh_m} := HINV. + by move=> [xc fx yc fy] [#] _; rewrite Sch_notdomE_hs; case: HINV. +have H /H {H}:= build_hpath_down mh y1 (ch + 1) x1 ch p v h _. ++ move=> p0 v0; rewrite -negP=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H} /# := HINV. + by move=> p' b' v' h' <*>; rewrite Sch_notrngE2_mh. ++ move=> ^ /build_hpathP + -> /=; rewrite !get_setE. + by case=> [<*>|/#]; move: HINV=> /hs_of_INV [] _ + H - /H {H} /#. +(* + by apply(ro_p_of_INV _ _ _ _ _ _ _ _ _ HINV). *) +split=>[]. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ [] ->//. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ []. ++ move=>l hmem i hi. + have[]_ _ h2 h3:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have[]sa sc[]:=h2 l hmem i hi. + have h1:=all_prefixes_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + smt(domE get_setE). +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +qed. + +lemma lemma2 hs ch PFm PFmi G1m G1mi G1mh G1mhi ro pi prefixes queries x1 x2 y1 y2 hx: + INV_CF_G1 hs ch PFm PFmi G1m G1mi G1mh G1mhi ro pi prefixes queries + => PFm.[(x1,x2)] = None + => G1m.[(x1,x2)] = None + => pi.[x2] = None + => hs.[hx] = Some (x2,Known) + => (forall f h, hs.[h] <> Some (y2,f)) + => INV_CF_G1 hs.[ch <- (y2,Known)] (ch + 1) + PFm.[(x1,x2) <- (y1,y2)] PFmi.[(y1,y2) <- (x1,x2)] + G1m.[(x1,x2) <- (y1,y2)] G1mi.[(y1,y2) <- (x1,x2)] + G1mh.[(x1,hx) <- (y1,ch)] G1mhi.[(y1,ch) <- (x1,hx)] + ro pi prefixes queries. +proof. +move=> HINV PFm_x1x2 G1m_x1x2 pi_x2 hs_hx y2_notrngE1_hs. +split. ++ by apply/hs_addh=> //=; case: HINV. ++ apply/inv_addm=> //; 1:by case: HINV. + case: {-1}(G1mi.[(y1,y2)]) (eq_refl G1mi.[(y1,y2)])=> [//|[xa xc]]. + + have /incli_of_INV @/incl + ^h - <- := HINV; 1: by rewrite h. + have /mi_mhi_of_INV [] H _ /H {H} [hx' fx' hy' fy'] [#] := HINV. + by rewrite y2_notrngE1_hs. ++ apply/inv_addm; 1:by case: HINV. + + have ^ /m_mh_of_INV Hm_mh /hs_of_INV Hhs := HINV. + by apply/(notin_m_notin_mh _ _ _ _ Hm_mh PFm_x1x2 hs_hx). + have ^ /mi_mhi_of_INV Hmi_mhi /hs_of_INV Hhs := HINV. + by apply/(ch_notdomE2_mh _ _ Hmi_mhi Hhs). ++ have ^ /hs_of_INV ^ Hhs [] Hhuniq _ _ /m_mh_of_INV := HINV. + move=> /m_mh_addh_addm /(_ hx x1 x2 ch y1 y2 Known Known Hhuniq hs_hx _) //. + exact/ch_notdomE_hs. ++ have ^ /hs_of_INV ^ Hhs [] Hhuniq _ _ /mi_mhi_of_INV := HINV. + move=> /mi_mhi_addh_addmi /(_ hx x1 x2 ch y1 y2 Known Known _ hs_hx _) //. + exact/ch_notdomE_hs. ++ by have /incl_of_INV/incl_addm ->:= HINV. ++ by have /incli_of_INV/incl_addm ->:= HINV. ++ split. + + move=> xa' hx' ya' hy'; rewrite get_setE; case: ((xa',hx') = (x1,hx))=> [[#] <*>> [#] <<*> /=|]. + + exists x2 Known y2 Known=> //=; rewrite !get_setE /=. + by have /hs_of_INV [] _ _ dom_hs /#:= HINV. + move=> xahx'_neq_x1hx; have /mh_of_INV [] Hmh _ _ /Hmh {Hmh} := HINV. + move=> [xc fx yc] [] /= [#] hs_hx' hs_hy'=> [[p v] [Hro Hpath]|<*> Gm_xa'xc]. + + exists xc fx yc Unknown=> /=; rewrite !get_setE hs_hx' hs_hy'. + rewrite (dom_hs_neq_ch hs xc fx _ hs_hx') /=; 1:by case: HINV. + rewrite (dom_hs_neq_ch hs yc Unknown _ hs_hy')/= ; 1:by case: HINV. + exists p v; rewrite Hro /=; apply/build_hpath_up/(notin_m_notin_mh _ _ _ _ _ PFm_x1x2 hs_hx). + + done. + by case: HINV. + exists xc Known yc Known=> //=; rewrite !get_setE; case: ((xa',xc) = (x1,x2))=> [/#|]. + rewrite Gm_xa'xc /= (dom_hs_neq_ch hs xc Known _ hs_hx') /=; 1:by case: HINV. + by rewrite (dom_hs_neq_ch hs yc Known _ hs_hy')/= ; 1:by case: HINV. + + move=> p xa b; have /mh_of_INV [] _ -> _ := HINV; split. + + move=> [v hi hf] [#] Hpath mh_vxahi; exists v hi hf. + rewrite get_setE; case: ((v +^ xa,hi) = (x1,hx))=> [[#] <*>|_]. + + move: mh_vxahi; have /m_mh_of_INV [] _ H /H {H}:= HINV. + by move=> [xc fx yc fy] [#]; rewrite hs_hx=> [#] <*>; rewrite PFm_x1x2. + rewrite mh_vxahi /=; apply/build_hpath_up=> //. + by apply/(notin_m_notin_mh _ _ _ _ _ PFm_x1x2 hs_hx); case: HINV. + move=> [v hi hf] [#]. + have no_path_to_hx: forall p0 v0, build_hpath G1mh p0 <> Some (v0,hx). + + have /pi_of_INV [] /(_ x2):= HINV; rewrite pi_x2 /=. + by move=> + p0 v0 - /(_ p0 v0) /negb_exists /(_ hx) /=; rewrite hs_hx. + have H /H {H} := build_hpath_down G1mh x1 hx y1 ch p v hi no_path_to_hx. + rewrite get_setE. case: ((v +^ xa,hi) = (x1,hx))=> [[#] <*>|_ Hpath Hextend]. + + by rewrite no_path_to_hx. + by exists v hi hf. + move=> p v p' v' h0. + have: forall p0 v0, build_hpath G1mh p0 <> Some (v0,hx). + + have /pi_of_INV [] /(_ x2):= HINV; rewrite pi_x2 /=. + by move=> + p0 v0 - /(_ p0 v0) /negb_exists /(_ hx) /=; rewrite hs_hx. + move=> ^ + /build_hpath_down H /H {H} - /build_hpath_down H + /H {H}. + by have /mh_of_INV [] _ _ /(_ p v p' v' h0):= HINV. +split=> c p v; have /pi_of_INV [] -> := HINV. +apply/exists_iff=> h /=; split=> [#]. ++ move=> /build_hpath_up /(_ x1 hx y1 ch _). + + by apply/(notin_m_notin_mh hs PFm x2 Known); case:HINV. + move=> -> /=; rewrite get_setE. + by have /hs_of_INV [] _ _ dom_hs ^ + /dom_hs /#:= HINV. +have no_path_to_hx: forall p0 v0, build_hpath G1mh p0 <> Some (v0,hx). ++ have /pi_of_INV [] /(_ x2):= HINV; rewrite pi_x2 /=. + by move=> + p0 v0 - /(_ p0 v0) /negb_exists /(_ hx) /=; rewrite hs_hx. +have H /H {H} := build_hpath_down G1mh x1 hx y1 ch p v h no_path_to_hx. +move=> ^ Hpath -> /=; rewrite get_setE; case: (h = ch)=> [<*> /= [#] <*>|//=]. +move: Hpath=> /build_hpathP [<*>|]. ++ by have /hs_of_INV [] _ + H - /H {H}:= HINV. ++ move=> p' b' v' h' <*> _; have /m_mh_of_INV [] _ H /H {H}:= HINV. + by move=> [xc fx yc fy] [#] _; have /hs_of_INV [] _ _ H /H {H}:= HINV. +(* + by apply(ro_p_of_INV _ _ _ _ _ _ _ _ _ HINV). *) +split=>[]. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ [] ->//. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ []. ++ move=>l hmem i hi. + have[]_ _ h2 _:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have[]sa sc[]:=h2 l hmem i hi. + have h1:=all_prefixes_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + smt(domE get_setE). +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +qed. + +lemma lemma2' hs ch PFm PFmi G1m G1mi G1mh G1mhi ro pi prefixes queries x1 x2 y1 y2 hx: + INV_CF_G1 hs ch PFm PFmi G1m G1mi G1mh G1mhi ro pi prefixes queries + => ! (y1,y2) \in PFm + => PFmi.[(x1,x2)] = None + => G1mi.[(x1,x2)] = None + => hs.[hx] = Some (x2,Known) + => (forall f h, hs.[h] <> Some (y2,f)) + => INV_CF_G1 hs.[ch <- (y2,Known)] (ch + 1) + PFm.[(y1,y2) <- (x1,x2)] PFmi.[(x1,x2) <- (y1,y2)] + G1m.[(y1,y2) <- (x1,x2)] G1mi.[(x1,x2) <- (y1,y2)] + G1mh.[(y1,ch) <- (x1,hx)] G1mhi.[(x1,hx) <- (y1,ch)] + ro pi prefixes queries. +proof. +move=> HINV hh PFmi_x1x2 G1mi_x1x2 hs_hx y2_notrngE1_hs. +split. ++ by apply/hs_addh=> //=; case: HINV. ++ apply/inv_addm=> //; 1:by case: HINV. + case: {-1}(G1m.[(y1,y2)]) (eq_refl G1m.[(y1,y2)])=> [//|[xa xc]]. + + have /incl_of_INV + ^h - <- := HINV; 1: by rewrite h. + have /m_mh_of_INV [] H _ /H {H} [hx' fx' hy' fy'] [#] := HINV. + by rewrite y2_notrngE1_hs. ++ apply/inv_addm; 1:by case: HINV. + + have ^ /m_mh_of_INV Hm_mh /hs_of_INV Hhs := HINV. + by apply/(ch_notdomE2_mh _ _ Hm_mh Hhs). + have ^ /mi_mhi_of_INV Hm_mh /hs_of_INV Hhs := HINV. + by apply/(notin_m_notin_mh _ _ _ _ Hm_mh PFmi_x1x2 hs_hx). ++ have ^ /hs_of_INV ^ Hhs [] Hhuniq _ _ /m_mh_of_INV := HINV. + move=> /mi_mhi_addh_addmi /(_ hx x1 x2 ch y1 y2 Known Known _ hs_hx _) //. + exact/ch_notdomE_hs. ++ have ^ /hs_of_INV ^ Hhs [] Hhuniq _ _ /mi_mhi_of_INV := HINV. + move=> /m_mh_addh_addm /(_ hx x1 x2 ch y1 y2 Known Known _ hs_hx _) //. + exact/ch_notdomE_hs. ++ by have /incl_of_INV/incl_addm ->:= HINV. ++ by have /incli_of_INV/incl_addm ->:= HINV. ++ split. + + move=> ya' hy' xa' hx'; rewrite get_setE; case: ((ya',hy') = (y1,ch))=> [[#] <*>> [#] <<*> /=|]. + + exists y2 Known x2 Known=> //=; rewrite !get_setE /=. + by have /hs_of_INV [] _ _ dom_hs /#:= HINV. + move=> yahy'_neq_y1ch; have /mh_of_INV [] Hmh _ _ /Hmh {Hmh} := HINV. + move=> [yc fy xc] [] /= [#] hs_hy' hs_hx'=> [[p v] [#] Hro Hpath|Gm_ya'yc <*>]. + + exists yc fy xc Unknown => /=; rewrite !get_setE hs_hx' hs_hy'. + rewrite (dom_hs_neq_ch hs yc fy _ hs_hy') /=; 1:by case: HINV. + rewrite (dom_hs_neq_ch hs xc Unknown _ hs_hx')/= ; 1:by case: HINV. + exists p v; rewrite Hro /=; apply/build_hpath_up=> //. + case: {-1}(G1mh.[(y1,ch)]) (eq_refl G1mh.[(y1,ch)])=> [//|[za zc]]. + have /m_mh_of_INV [] _ H /H {H} [? ? ? ?] [#]:= HINV. + by have /hs_of_INV [] _ _ H /H {H} := HINV. + exists yc Known xc Known=> //=; rewrite !get_setE; case: ((ya',yc) = (y1,y2))=> [/#|]. + rewrite Gm_ya'yc /= (dom_hs_neq_ch hs yc Known _ hs_hy') /=; 1:by case: HINV. + by rewrite (dom_hs_neq_ch hs xc Known _ hs_hx')/= ; 1:by case: HINV. + + move=> p ya b; have /mh_of_INV [] _ -> _ := HINV. + apply/exists_iff=> v /=; apply/exists_iff=> hx' /=; apply/exists_iff=> hy' /=. + split=> [#]. + + move=> /(@build_hpath_up _ y1 ch x1 hx) /(_ _). + + apply/(@notin_hs_notdomE2_mh hs PFm)/(ch_notdomE_hs); by case: HINV. + move=> -> /=; rewrite get_setE /=; case: (hx' = ch)=> <*> //. + have /m_mh_of_INV [] _ H /H {H} [xc fx yc fy] [#] := HINV. + by have /hs_of_INV [] _ _ H /H {H} := HINV. + have no_path_to_ch: forall p0 v0, build_hpath G1mh p0 <> Some (v0,ch). + + move=> p0 v0; elim/last_ind: p0. + + by have /hs_of_INV [] /# := HINV. + move=> p0 b0 _; rewrite build_hpath_prefix. + apply/negb_exists=> b' /=; apply/negb_exists=> h' /=; apply/negb_and=> /=; right. + rewrite -negP; have /mh_of_INV [] H _ _ /H {H} [? ? ? ?] [#] _ := HINV. + by have /hs_of_INV [] _ _ H /H {H} := HINV. + have H /H {H} := build_hpath_down G1mh y1 ch x1 hx p v hx' no_path_to_ch. + rewrite get_setE. case: ((v +^ ya,hx') = (y1,ch))=> [[#] <*>|_ Hpath Hextend //=]. + by rewrite no_path_to_ch. + move=> p v p' v' h0. + have: forall p0 v0, build_hpath G1mh p0 <> Some (v0,ch). + + move=> p0 v0; elim/last_ind: p0. + + by have /hs_of_INV [] /# := HINV. + move=> p0 b0 _; rewrite build_hpath_prefix. + apply/negb_exists=> b' /=; apply/negb_exists=> h' /=; apply/negb_and=> /=; right. + rewrite -negP; have /mh_of_INV [] H _ _ /H {H} [? ? ? ?] [#] _ := HINV. + by have /hs_of_INV [] _ _ H /H {H} := HINV. + move=> ^ + /build_hpath_down H /H {H} - /build_hpath_down H + /H {H}. + by have /mh_of_INV [] _ _ /(_ p v p' v' h0):= HINV. +split=> c p v; have /pi_of_INV [] -> := HINV. +apply/exists_iff=> h /=; split=> [#]. ++ move=> /build_hpath_up /(_ y1 ch x1 hx _). + + have ^ /m_mh_of_INV [] _ H /hs_of_INV [] _ _ H' := HINV. + case: {-1}(G1mh.[(y1,ch)]) (eq_refl (G1mh.[(y1,ch)]))=> [//|]. + by move=> [za zc] /H [? ? ? ?] [#] /H'. + move=> -> /=; rewrite get_setE. + by have /hs_of_INV [] _ _ dom_hs ^ + /dom_hs /#:= HINV. +have no_path_to_ch: forall p0 v0, build_hpath G1mh p0 <> Some (v0,ch). ++ move=> p0 v0; elim/last_ind: p0. + + by have /hs_of_INV [] /# := HINV. + move=> p0 b0 _; rewrite build_hpath_prefix. + apply/negb_exists=> b' /=; apply/negb_exists=> h' /=; apply/negb_and=> /=; right. + rewrite -negP; have /mh_of_INV [] H _ _ /H {H} [? ? ? ?] [#] _ := HINV. + by have /hs_of_INV [] _ _ H /H {H} := HINV. ++ have H /H {H} := build_hpath_down G1mh y1 ch x1 hx p v h no_path_to_ch. + move=> ^ Hpath -> /=; rewrite get_setE; case: (h = ch)=> [<*> /= [#] <*>|//=]. + move: Hpath=> /build_hpathP [<*>|]. + + by have /hs_of_INV [] _ + H - /H {H}:= HINV. + move=> p' b' v' h' <*> _; have /m_mh_of_INV [] _ H /H {H}:= HINV. + by move=> [xc fx yc fy] [#] _; have /hs_of_INV [] _ _ H /H {H}:= HINV. +(* + by apply(ro_p_of_INV _ _ _ _ _ _ _ _ _ HINV). *) +split=>[]. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ []. ++ by case:HINV=>_ _ _ _ _ _ _ _ (* _ *) [] _ []. ++ move=>l hmem i hi. + have[]_ _ h2 _:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have[]sa sc[]:=h2 l hmem i hi. + have h1:=all_prefixes_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + smt(domE get_setE). +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +by case:HINV=>_ _ _ _ _ _ _ _ _ []. +qed. + +lemma lemma3 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries xa xc hx ya yc hy p b: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries + => Pm.[(xa,xc)] = Some (ya,yc) + => Gm.[(xa,xc)] = None + => mh.[(xa,hx)] = Some (ya,hy) + => hs.[hx] = Some (xc,Known) + => hs.[hy] = Some (yc,Unknown) + => pi.[xc] = Some (p,b) + => INV_CF_G1 hs.[hy <- (yc,Known)] ch + Pm Pmi + Gm.[(xa,xc) <- (ya,yc)] Gmi.[(ya,yc) <- (xa,xc)] + mh mhi + ro pi.[yc <- (rcons p (b +^ xa),ya)] prefixes queries. +proof. +move=> HINV Pm_xaxc Gm_xaxc mh_xahx hs_hx hs_hy pi_xc. +split. ++ have /hs_of_INV /hs_updh /(_ Unknown) H := HINV; apply/H=> {H} //. + by rewrite -negP=> <*>; move: hs_hy; have /hs_of_INV [] _ -> := HINV. ++ apply/inv_addm=> //; 1:by case: HINV. + case: {-1}(Gmi.[(ya,yc)]) (eq_refl Gmi.[(ya,yc)])=> [//|[xa' xc']]. + have /incli_of_INV + ^h - <- := HINV; 1:by rewrite h. + move: Pm_xaxc; have [] -> -> /= := inv_mh_inv_Pm hs Pm Pmi mh mhi _ _ _; first 3 by case: HINV. + rewrite -negP=> [#] <<*>. + move: h; have /invG_of_INV [] <- := HINV. + by rewrite Gm_xaxc. ++ by case: HINV. ++ by apply/(m_mh_updh Unknown)=> //; case: HINV. ++ by apply/(m_mh_updh Unknown)=> //; case: HINV. ++ move=> [za zc]; rewrite get_setE; case: ((za,zc) = (xa,xc))=> // _. + by have /incl_of_INV H /H {H}:= HINV. ++ move: mh_xahx; have /inv_of_INV [] H /H {H}:= HINV. + have /mi_mhi_of_INV [] _ H /H {H} [xct fxt yct fyt] [#] := HINV. + rewrite hs_hx hs_hy=> /= [#] 2!<<- [#] 2!<<- Pmi_yayc. + move=> [za zc]; rewrite get_setE; case: ((za,zc) = (ya,yc))=> // _. + by have /incli_of_INV H /H {H}:= HINV. ++ split; last 2 by have /mh_of_INV [] _:= HINV. + move=> xa' hx' ya' hy'; case: ((xa',hx') = (xa,hx))=> [[#] <*>|]. + + rewrite mh_xahx=> /= [#] <<*>; rewrite !get_setE /=. + case: (hx = hy)=> [<*>|_]; first by move: hs_hx; rewrite hs_hy. + by exists xc Known yc Known; rewrite get_setE. + move=> Hxahx' mh_xahx'. + have ^path_to_hy: build_hpath mh (rcons p (b +^ xa)) = Some (ya,hy). + + apply/build_hpath_prefix; exists b hx. + rewrite xorwA xorwK xorwC xorw0 mh_xahx /=. + move: pi_xc; have /pi_of_INV [] -> [h] [#] := HINV. + by have /hs_of_INV [] H _ _ + /H {H} /(_ _ _ hs_hx _) := HINV. + have /mh_of_INV [] /(_ _ _ _ _ mh_xahx') + ro_def H /H {H} unique_path_to_hy := HINV. + move=> [xc' fx' yc' fy'] /= [#]. + case: (hy' = hy)=> [<*> hs_hx'|Hhy']. + + rewrite hs_hy=> /= [#] <<*> /= [p' b'] [#] ro_pbxa' path_hx'. + have:= unique_path_to_hy (rcons p' (b' +^ xa')) ya' _. + + by apply/build_hpath_prefix; exists b' hx'; rewrite xorwA xorwK xorwC xorw0. + move=> [#] ^/rconsIs + /rconssI - <<*>. + by move: mh_xahx' Hxahx' mh_xahx; have /inv_of_INV [] ^ + -> - -> -> /= -> := HINV. + rewrite (@get_set_neqE _ _ hy' _ Hhy')=> /= hs_hx' ^ hs_hy' -> Hite. + exists xc' (if hx' = hy then Known else fx') yc' fy'. + rewrite (@get_setE Gm) (_: (xa',xc') <> (xa,xc)) /=. + + move: Hxahx'=> /=; case: (xa' = xa)=> [<*> /=|//]. + by apply/contra=> <*>; have /hs_of_INV [] + _ _ - /(_ _ _ _ _ hs_hx' hs_hx _) := HINV. + rewrite get_setE; case: (hx' = hy)=> /= [<*>|//]. + move: hs_hx'; rewrite hs_hy=> /= [#] <<*> /=. + by move: Hite=> /= [#]; case: fy' hs_hy'=> //= _ ->. ++ split=> c p' b'; rewrite !get_setE; case: (yc = c)=> [<<*> /=|yc_neq_c]; last first. + + rewrite (@eq_sym c) yc_neq_c /=; have /pi_of_INV [] -> := HINV. + apply/exists_iff=> h /=; rewrite get_setE; case: (h = hy)=> [<*> /=|//=]. + by rewrite yc_neq_c hs_hy /=. + split=> [[#] <<*>|]. + + exists hy; rewrite get_setE /=; apply/build_hpath_prefix. + exists b hx; rewrite xorwA xorwK xorwC xorw0 mh_xahx /=. + move: pi_xc; have /pi_of_INV [] -> [h] [#] + hs_h:= HINV. + by have /hs_of_INV [] + _ _ - /(_ _ _ _ _ hs_hx hs_h _) := HINV. + move=> [h]; rewrite get_setE; case: (h = hy)=> [<*> /=|]; last first. + + by have /hs_of_INV [] H _ _ + [#] _ /H {H} /(_ _ _ hs_hy _) // <*> := HINV. + have /mh_of_INV [] _ _ /(_ p' b') H /H {H} /(_ (rcons p (b +^ xa)) ya _) //:= HINV. + apply/build_hpath_prefix; exists b hx; rewrite xorwA xorwK xorwC xorw0 mh_xahx /=. + move: pi_xc; have /pi_of_INV [] -> [h] [#] + hs_h:= HINV. + by have /hs_of_INV [] + _ _ - /(_ _ _ _ _ hs_hx hs_h _) := HINV. +(* + by apply(ro_p_of_INV _ _ _ _ _ _ _ _ _ HINV). *) ++ by case:HINV. +qed. + + +lemma m_mh_None hs0 PFm G1mh hx2 x2 k x1: + m_mh hs0 PFm G1mh => + hs0.[hx2] = Some (x2, k) => + PFm.[(x1, x2)] = None => + G1mh.[(x1,hx2)] = None. +proof. + move=> [] HP /(_ x1 hx2) + Hhx2;case (G1mh.[(x1, hx2)]) => //. + by move=> -[ya hy] /(_ ya hy) /= [] ????; rewrite Hhx2 => /= [#] <- _ _ ->. +qed. + +lemma build_hpath_None (G1mh:hsmap) p: + foldl (step_hpath G1mh) None p = None. +proof. by elim:p. qed. + +lemma build_hpath_upd_ch ha ch mh xa ya p v hx: + 0 <> ch => ha <> ch => (forall xa xb ha hb, mh.[(xa,ha)] = Some(xb, hb) => ha <> ch /\ hb <> ch) => + build_hpath mh.[(xa, ha) <- (ya, ch)] p = Some (v, hx) => + if hx = ch then + (exists p0 x, build_hpath mh p0 = Some (x, ha) /\ p = rcons p0 (x +^ xa) /\ v = ya) + else + build_hpath mh p = Some (v, hx). +proof. + move=> Hch0 Hha Hch. + elim/last_ind: p v hx=> /=. + + by move=> v hx;rewrite /build_hpath /= => -[!<<-] //; rewrite Hch0. + move=> p x Hrec v hx /build_hpath_prefix [v' h' [/Hrec{Hrec}]]. + rewrite get_setE /=;case (h' = ch) => [->> | ]. + + by rewrite (@eq_sym ch) Hha /= => _ /Hch. + case (v' +^ x = xa /\ h' = ha) => [[!<<-] /= ?? [!->>] /=| ]. + + by exists p v'; rewrite xorwA xorwK xorwC xorw0. + case (hx = ch)=> [->> |??? Hbu Hg]. + + by move=> ??? /= /Hch. + by rewrite build_hpath_prefix;exists v' h';smt(). +qed. + +lemma build_hpath_up_None (G1mh:hsmap) bi1 bi2 bi p: + G1mh.[bi1] = None => + build_hpath G1mh p = Some bi => + build_hpath G1mh.[bi1 <- bi2] p = Some bi. +proof. + rewrite /build_hpath;move=> Hbi1. + elim: p (Some (b0,0)) => //= b p Hrec obi. + rewrite {2 4}/step_hpath /=;case: obi => //= [ | bi'];1:by apply Hrec. + rewrite get_setE. case ((bi'.`1 +^ b, bi'.`2) = bi1) => [-> | _];2:by apply Hrec. + by rewrite Hbi1 build_hpath_None. +qed. + +(* +lemma build_hpath_down_None h ch mh xa ha ya a p: + h <> ch => ha <> ch => + (forall ya, mh.[(ya,ch)] = None) => + build_hpath mh.[(xa,ha) <- (ya,ch)] p = Some (a,h) => + build_hpath mh p = Some (a,h). +proof. + move=> Hh Hha Hmh;rewrite /build_hpath;move: (Some (b0, 0)). + elim: p => //= b p Hrec [ | bi] /=;rewrite {2 4}/step_hpath /= ?build_hpath_None //. + rewrite oget_some get_setE;case ((bi.`1 +^ b, bi.`2) = (xa, ha)) => _;2:by apply Hrec. + move=> {Hrec};case: p=> /= [[_ ->>]| b' p];1: by move:Hh. + by rewrite {2}/step_hpath /= oget_some /= get_setE_neq /= ?Hha // Hmh build_hpath_None. +qed. +*) + +lemma build_hpath_upd_ch_iff ha ch mh xa ya p v hx: + mh.[(xa,ha)] = None => + 0 <> ch => ha <> ch => (forall xa xb ha hb, mh.[(xa,ha)] = Some(xb, hb) => ha <> ch /\ hb <> ch) => + build_hpath mh.[(xa, ha) <- (ya, ch)] p = Some (v, hx) <=> + if hx = ch then + (exists p0 x, build_hpath mh p0 = Some (x, ha) /\ p = rcons p0 (x +^ xa) /\ v = ya) + else + build_hpath mh p = Some (v, hx). +proof. + move=> Ha Hch0 Hha Hch;split;1: by apply build_hpath_upd_ch. + case (hx = ch);2: by move=> ?;apply build_hpath_up_None. + move=> ->> [p0 x [? [!->>]]]. + rewrite build_hpath_prefix;exists x ha. + by rewrite xorwA xorwK xorwC xorw0 get_set_sameE /=;apply build_hpath_up_None. +qed. + + +lemma lemma4 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries i p sa sc h f: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries +=> 0 <= i < List.size p +=> take (i + 1) p \in prefixes +=> prefixes.[take i p] = Some (sa,sc) +=> build_hpath mh (take i p) = Some (sa,h) +=> ro.[take (i+1) p] = Some (oget prefixes.[take (i+1) p]).`1 +=> hs.[h] = Some (sc, f) +=> (sa +^ nth witness p i, h) \in mh. +proof. +move=>inv0 hi take_i1_p_in_prefixes prefixes_sa_sc build_hpath_i_p ro_prefix hs_h_sc_f. +have[]_ _ m_prefix _:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. +have[]b1 c1[]:=m_prefix _ take_i1_p_in_prefixes i _;1:smt(size_take). +rewrite!take_take!minrE //= (: i <= i + 1) 1:/# nth_take 1,2:/# prefixes_sa_sc/==>[][<-<-]{b1 c1}Pm_prefix. +have[]hh1 hh2 hh3:=mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. +move:ro_prefix;have{1}->:=(take_nth witness i p);1:smt(size_take);move=>h1. +have:=hh2 (take i p) (nth witness p i) (oget prefixes.[take (i + 1) p]).`1. +rewrite h1/==>[][] v hx hy;rewrite build_hpath_i_p/==>[][][?<-];smt(domE). +qed. + +(* we should do a lemma to have the equivalence *) + + +equiv eq_fi (D <: DISTINGUISHER {-PF, -RO, -G1}): DPRestr(PF).fi ~ DPRestr(G1(DRestr(D)).S).fi: + !G1.bcol{2} + /\ !G1.bext{2} + /\ ={arg} /\ ={glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2} + ==> if G1.bcol{2} \/ G1.bext{2} + then ([] \in C.queries{1}) /\ ([] \in C.queries{2}) + else ={res} /\ ={glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2}. +proof. +proc;sp;if;auto. +call(: !G1.bcol{2} + /\ !G1.bext{2} + /\ ={x} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2} + ==> !G1.bcol{2} + => !G1.bext{2} + => ={res} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2});auto. +exists* FRO.m{2}, G1.chandle{2}, PF.m{1}, PF.mi{1}, + G1.m{2}, G1.mi{2}, G1.mh{2}, G1.mhi{2}, + F.RO.m{2}, G1.paths{2}, x{2}, Redo.prefixes{1}, C.queries{2}. +elim* => hs ch Pm Pmi Gm Gmi mh mhi ro pi [xa xc] prefixes queries. +case @[ambient]: + {-1}(INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries) + (eq_refl (INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries)); last first. ++ by move=> inv0; exfalso=> ? ? [#] <<*>; rewrite inv0. +move=> /eqT inv0; proc. +case @[ambient]: {-1}(Pmi.[(xa,xc)]) (eq_refl Pmi.[(xa,xc)])=> [Pmi_xaxc|[ya yc] Pmi_xaxc]. ++ have /incli_of_INV /(_ (xa,xc)) := inv0; rewrite Pmi_xaxc /=. + case: {-1}(Gmi.[(xa,xc)]) (eq_refl Gmi.[(xa,xc)])=> //= Gmi_xaxc. + rcondt{1} 1; 1:by auto=> &hr [#] <<*>; rewrite domE Pmi_xaxc. + rcondt{2} 1; 1:by auto=> &hr [#] <<*>; rewrite domE Gmi_xaxc. + case @[ambient]: {-1}(getflag hs xc) (eq_refl (getflag hs xc)). + + move=> /getflagP_none xc_notrngE1_hs. + rcondt{2} 2. + + auto=> &hr [#] <<*> _ _ _; rewrite rngE /= negb_exists=> h /=. + by rewrite xc_notrngE1_hs. + rcondf{2} 8. + + auto=> &hr [#] !<<- _ _ ->> _ /= _ _ _ _. + rewrite negb_and domE; left. + rewrite (@huniq_hinvK_h ch) 3:oget_some /=. + + by apply/huniq_addh=> //; have /hs_of_INV [] := inv0. + + by rewrite get_setE. + apply/(@notin_m_notin_mh hs.[ch <- (xc,Known)] Pmi _ _ xc ch Known)=> //. + + by apply/m_mh_addh=> //; case: inv0. + by rewrite get_setE. + auto=> ? ? [#] !<<- -> -> ->> _ /= ya ? /= yc ? /=. + case: (hinvP (hs.[ch <- (xc,Known)]) yc)=> [_|-> //] yc_notrngE1_hs_addh _ _. + rewrite get_setE /= (@huniq_hinvK_h ch) 3:oget_some /=. + + by apply/huniq_addh=> //; have /hs_of_INV [] := inv0. + + by rewrite get_setE. + apply/(@lemma1' hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries xa xc ya yc inv0 _ _ Pmi_xaxc Gmi_xaxc)=> //;first last. + + rewrite -negP=> <*>; move: yc_notrngE1_hs_addh => /=. + apply/negb_forall=> /=; exists ch; apply/negb_forall=> /=; exists Known. + by rewrite get_setE. + + move=> f h; move: (yc_notrngE1_hs_addh h f); rewrite get_setE. + case: (h = ch)=> <*> //= _; rewrite -negP. + by have /hs_of_INV [] _ _ H /H {H} := inv0. + + rewrite domE/=;have[]h1 h2:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + have h1':=h1 ya yc. + have :Pm.[(ya, yc)] <> None => exists (hx : handle) (fx : flag), hs.[hx] = Some (yc, fx). + + move=> y_in_Pm; move: (h1' (oget Pm.[(ya,yc)]).`1 (oget Pm.[(ya,yc)]).`2 _). + + by move: y_in_Pm; case: (Pm.[(ya,yc)])=> - //= []. + by move=> [hx fx hy fy] [#] h _ _; exists hx fx. + case(Pm.[(ya, yc)] = None)=>//=h; + rewrite negb_exists/==>a;rewrite negb_exists/==>b. + have:=yc_notrngE1_hs_addh a b;rewrite get_setE;case(a=ch)=>//=hach. + case(xc=yc)=>[/#|]hxyc. + have[]_ _ help:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + by have/#:=help (yc,b) a. + have /hs_of_INV [] Hhuniq _ _ [] /(getflagP_some _ _ _ Hhuniq):= inv0. + + move=> x2_is_U; conseq (_: _ ==> G1.bext{2})=> //. + by auto=> ? ? [#] !<<- _ -> ->>_ /=; rewrite x2_is_U. + move=> x2_is_K; have:=x2_is_K; rewrite rngE=> -[hx2] hs_hx2. + rcondf{2} 2; 1:by auto=> &hr [#] <*> /=; rewrite x2_is_K. + rcondf{2} 6. + + auto=> &hr [#] !<<- _ _ ->> _. + rewrite (@huniq_hinvK_h hx2) // oget_some /= => _ _ _ _. + rewrite negb_and domE /=; left. + by apply/(@notin_m_notin_mh hs Pmi _ _ xc _ Known)=> //; case: inv0. + auto => ? ? [#] !<<- -> -> ->> _. + rewrite (@huniq_hinvK_h hx2) // oget_some /= => y1 ? /= y2 ? /=. + case: (hinvP hs y2)=> [_ y2_notrngE1_hs _ _|/#]. + rewrite get_setE /=. + apply/lemma2'=> //. + + rewrite domE/=;have[]h1 _:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + have h1':=h1 y1 y2. + have :Pm.[(y1, y2)] <> None => exists (hx : handle) (fx : flag), hs.[hx] = Some (y2, fx). + + move=> y_in_Pm; move: (h1' (oget Pm.[(y1,y2)]).`1 (oget Pm.[(y1,y2)]).`2 _). + + by move: y_in_Pm; case: (Pm.[(y1,y2)])=> - //= []. + by move=> [hx fx hy fy] [#] h _ _; exists hx fx. + case(Pm.[(y1, y2)] = None)=>//=h; + rewrite negb_exists/==>a;rewrite negb_exists/==>b. + exact(y2_notrngE1_hs). + move=> f h; exact/y2_notrngE1_hs. +rcondf{1} 1; 1:by auto=> &hr [#] <<*>; rewrite domE Pmi_xaxc. +case @[ambient]: {-1}(Gmi.[(xa,xc)]) (eq_refl Gmi.[(xa,xc)])=> [|[ya' yc']] Gmi_xaxc. ++ rcondt{2} 1; 1:by auto=> &hr [#] <<*>; rewrite domE Gmi_xaxc. + conseq (_: _ ==> G1.bext{2})=> //. + auto=> &1 &2 [#] !<<- _ -> ->> _ />. + rewrite !rngE /=; have ->: exists hx, hs.[hx] = Some (xc,Unknown). + + move: Pmi_xaxc; have /mi_mhi_of_INV [] H _ /H {H} := inv0. + move=> [hx fx hy fy] [#] hs_hx hs_hy. + have ^/inv_of_INV [] <- /mh_of_INV [] H _ _ /H {H} := inv0. + move=> [? ? ? ?] [#]; rewrite hs_hx hs_hy=> /= [#] <<*> [#] <<*>. + case: fx hs_hx=> hs_hx /= => [_|[#]]; first by exists hx. + by have /invG_of_INV [] -> := inv0; rewrite Gmi_xaxc. + smt (@Block.DBlock @Capacity.DCapacity). +have:=Gmi_xaxc. +have /incli_of_INV <- := inv0; 1:by rewrite Gmi_xaxc. +rewrite Pmi_xaxc=> /= [#] <<*>. +rcondf{2} 1; 1:by auto=> &hr [#] <<*>; rewrite domE Gmi_xaxc. +by auto=> &1 &2 /#. ++ by move=> /> &1 &2 _ _ /m_p_of_INV []; smt(domE). +by move=> /> &1 &2 -> ->. +qed. + + +equiv eq_f (D <: DISTINGUISHER {-PF, -RO, -G1}): DPRestr(PF).f ~ DPRestr(G1(DRestr(D)).S).f: + !G1.bcol{2} + /\ !G1.bext{2} + /\ ={x} /\ ={glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2} + ==> if G1.bcol{2} \/ G1.bext{2} + then ([] \in C.queries{2}) + else ={res} /\ ={glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2}. +proof. +proc;sp;if;auto. +call(: !G1.bcol{2} + /\ !G1.bext{2} + /\ ={arg} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2} + ==> if G1.bcol{2} \/ G1.bext{2} + then ([] \in C.queries{2}) + else ={res} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2});auto. + exists * FRO.m{2}, G1.chandle{2}, + PF.m{1}, PF.mi{1}, + G1.m{2}, G1.mi{2}, G1.mh{2}, G1.mhi{2}, + F.RO.m{2}, G1.paths{2}, Redo.prefixes{1}, C.queries{2}, + x{2}. + elim * => hs0 ch0 PFm PFmi G1m G1mi G1mh G1mhi ro0 pi0 pref queries [] x1 x2. + (* poor man's extraction of a fact from a precondition *) + case @[ambient]: {-1}(INV_CF_G1 hs0 ch0 PFm PFmi G1m G1mi G1mh G1mhi ro0 pi0 pref queries) + (eq_refl (INV_CF_G1 hs0 ch0 PFm PFmi G1m G1mi G1mh G1mhi ro0 pi0 pref queries)); last first. + + by move=> h; exfalso=> &1 &2 [#] <*>; rewrite h. + move=> /eqT inv0; proc; case @[ambient] {-1}(PFm.[(x1,x2)]) (eq_refl (PFm.[(x1,x2)])). + + move=> PFm_x1x2. + have /incl_of_INV /(notin_m_notin_Gm _ _ (x1,x2)) /(_ _) // Gm_x1x2 := inv0. + rcondt{1} 1; 1:by move=> //= &1; skip=> &2 [#] <*>; rewrite domE PFm_x1x2. + rcondt{2} 1; 1:by move=> //= &1; skip=> &2 [#] <*>; rewrite domE Gm_x1x2. + case @[ambient]: {-1}(pi0.[x2]) (eq_refl (pi0.[x2])). + + move=> x2_in_pi; rcondf{2} 1. + + by move=> //= &1; skip=> &2 [#] <*>; rewrite domE x2_in_pi. + rcondf{2} 8. + + by move=> //= &1; auto=> &2 [#] !<<-; rewrite !domE x2_in_pi. + seq 2 2: ( hs0 = FRO.m{2} + /\ ch0 = G1.chandle{2} + /\ PFm = PF.m{1} + /\ PFmi = PF.mi{1} + /\ G1m = G1.m{2} + /\ G1mi = G1.mi{2} + /\ G1mh = G1.mh{2} + /\ G1mhi = G1.mhi{2} + /\ ro0 = F.RO.m{2} + /\ pi0 = G1.paths{2} + /\ pref = Redo.prefixes{1} + /\ queries = C.queries{2} + /\ (x1,x2) = x{2} + /\ !G1.bcol{2} + /\ !G1.bext{2} + /\ ={x, y1, y2} + /\ INV_CF_G1 hs0 ch0 PFm PFmi G1m G1mi G1mh G1mhi ro0 pi0 pref queries). + + by auto. + case @[ambient]: {-1}(getflag hs0 x2) (eq_refl (getflag hs0 x2)). + + rewrite getflagP_none => x2f_notrngE_hs0; rcondt{2} 3. + + move=> &1; auto=> &2 /> _ _ _; rewrite rngE /= negb_exists /=. + exact/(@x2f_notrngE_hs0 Known). + rcondf{2} 6. + + move=> &1; auto=> &2 />. + have ->: hinvK FRO.m{2}.[G1.chandle{2} <- (x2,Known)] x2 = Some G1.chandle{2}. + + rewrite (@huniq_hinvK_h G1.chandle{2} FRO.m{2}.[G1.chandle{2} <- (x2,Known)] x2) //. + + move=> hx hy [] xc xf [] yc yf /=. + rewrite !get_setE; case: (hx = G1.chandle{2}); case: (hy = G1.chandle{2})=> //=. + + by move=> _ + [#] - <*>; have:= (x2f_notrngE_hs0 yf hy). + + by move=> + _ + [#] - <*>; have:= (x2f_notrngE_hs0 xf hx). + by move=> _ _; have /hs_of_INV [] + _ _ - /(_ hx hy (xc,xf) (yc,yf)) := inv0. + by rewrite !get_setE. + rewrite oget_some=> _ _ _. + have -> //: (x1,G1.chandle{2}) \notin G1.mh{2}. + rewrite domE /=; case: {-1}(G1.mh.[(x1,G1.chandle)]{2}) (eq_refl (G1.mh.[(x1,G1.chandle)]{2}))=> //= -[xa xh]; rewrite -negP. + have ^/m_mh_of_INV [] _ + /hs_of_INV [] _ _ h_handles := inv0. + by move=> /(_ x1 G1.chandle{2} xa xh) h /h [] xc xf yc yf [#] /h_handles. + case: (x2 <> y2{2} /\ (forall f h, hs0.[h] <> Some (y2{2},f))). + + auto=> &1 &2 [#] !<<- -> -> !->> /= _ x2_neq_y2 y2_notin_hs. + rewrite get_setE /=. + rewrite (@huniq_hinvK_h ch0 hs0.[ch0 <- (x2,Known)] x2); 2:by rewrite get_setE. + + move=> @/huniq h1 h2 [c1 f1] [c2 f2]; rewrite !get_setE /=. + case: (h1 = ch0); case: (h2 = ch0)=> //=. + + by move=> _ + [#] - <*>; move: (x2f_notrngE_hs0 f2 h2). + + by move=> + _ + [#] <*> - <*>; move: (x2f_notrngE_hs0 f1 h1). + have /hs_of_INV [] + _ _ _ _ - h := inv0. + by apply/h; rewrite get_setE. + rewrite !oget_some;rewrite domE;have[]_ -> _ _ _ /=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + smt(lemma1). + conseq (_: _ ==> G1.bcol{2})=> //=. + + by auto=> &1 &2 [#] !<<- bad1 bad2 -> _ ->> !<<- _ /=/>; + rewrite domE;have[]_ ->_ _ _/=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + auto=> &1 &2 [#] !<<- -> _ ->> !<<- _ /=/>. + case: (hinvP hs0.[ch0 <- (x2,Known)] y2{1})=> //= -> /=. + move=> hs0_spec; split=> [|f]. + + by have:= hs0_spec ch0 Known; rewrite get_setE. + move=> h; have:= hs0_spec h f; rewrite get_setE; case: (h = ch0)=> [<*>|//=]. + by move=> _; rewrite -negP; have /hs_of_INV [] _ _ H /H {H}:= inv0. + case; rewrite getflagP_some; 1,3:by have /hs_of_INV []:= inv0. + + move=> x2_is_U; conseq (_: G1.bext{2})=> //=; auto=> &1 &2 /> _ _ hinv0 . + by rewrite domE;have[]_ -> _ _ _/=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + move=> x2_is_K; rcondf{2} 3; 1:by move=> &1; auto. + have:= x2_is_K; rewrite rngE=> - [hx] hs0_hx. + seq 0 3: ( hs0 = FRO.m{2} + /\ ch0 = G1.chandle{2} + /\ PFm = PF.m{1} + /\ PFmi = PF.mi{1} + /\ G1m = G1.m{2} + /\ G1mi = G1.mi{2} + /\ G1mh = G1.mh{2} + /\ G1mhi = G1.mhi{2} + /\ ro0 = F.RO.m{2} + /\ pi0 = G1.paths{2} + /\ pref = Redo.prefixes{1} + /\ queries = C.queries{2} + /\ (x1,x2) = x{2} + /\ !G1.bcol{2} + /\ !G1.bext{2} + /\ ={x,y1,y2} + /\ y{2} = (y1,y2){2} + /\ hx2{2} = hx + /\ INV_CF_G1 hs0 ch0 PFm PFmi G1m G1mi G1mh G1mhi ro0 pi0 pref queries). + (* TODO : reduce the example to reproduce the problem with : auto=> &1 &2 /> *) + + auto=> &1 &2 [#] 13!<<- 2!-> 3!->> HINV0 /=;split. + + move: x2_is_K; rewrite rngE /= => -[hx2] hs_hx2. + rewrite negb_exists /==> h; rewrite -negP=> hs_h. + have /hs_of_INV [] Hhuniq _ _ := inv0. + by move: (Hhuniq _ _ _ _ hs_hx2 hs_h)=> ht; move: ht hs_h=> /= <*>; rewrite hs_hx2. + rewrite (@huniq_hinvK_h hx hs0 x2) //. + by have /hs_of_INV [] := inv0. + have x1hx_notin_G1m: (x1,hx) \notin G1mh. + + rewrite domE; case: {-1}(G1mh.[(x1,hx)]) (eq_refl G1mh.[(x1,hx)])=> //=. + move=> [mhx1 mhx2]; rewrite -negP=> h. + have /m_mh_of_INV [] _ hg := inv0. + have [xa xh ya yh] := hg _ _ _ _ h. + by rewrite hs0_hx=> [#] <*>; rewrite PFm_x1x2. + rcondf{2} 1. + + by move=> &m; auto=> //= &hr [#] <*>; rewrite x1hx_notin_G1m. + auto=> &1 &2 [#] !<<- -> -> !->> _ /=. + rewrite domE;have[]_ -> _ _ _ /=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + case(hinv hs0 y2{2} = None)=>//=h; + rewrite get_setE /=;smt(lemma2 hinvP). + move=> [p0 v0] pi_x2; have:=pi_x2. + have /pi_of_INV [] -> [hx2] [#] Hpath hs_hx2:= inv0. + rcondt{2} 1. by move=> &m; auto=> &hr [#] !<<- _ _ ->> /= _; rewrite domE pi_x2. + rcondf{2} 6. + + auto; inline *; auto=> &hr [#] !<<- _ _ !->> _ /= _ _ _ _ /=. + by rewrite rngE; exists hx2. + rcondf{2} 7. + + auto; inline *; auto=> &hr [#] !<<- _ _ !->> _ /= _ _ _ _ /=. + rewrite negb_and; left; rewrite (@huniq_hinvK_h hx2 hs0 x2) // 2:oget_some. + + by have /hs_of_INV []:= inv0. + rewrite domE; case: {-1}(G1mh.[(x1,hx2)]) (eq_refl (G1mh.[(x1,hx2)]))=> [//=|[xa xc] G1mh_x1hx2]. + have /m_mh_of_INV [] _ /(_ _ _ _ _ G1mh_x1hx2) [xc0 xf0 yc0 yf0] := inv0. + by move=> [#]; rewrite hs_hx2=> [#] !<<-; rewrite PFm_x1x2. + rcondt{2} 15. + + auto; inline *; auto=> &hr [#] !<<- _ _ !->> _ /= _ _ _ _ /=. + by rewrite domE pi_x2. + inline F.RO.get. rcondt{2} 4. + + auto=> &hr [#] !<<- _ _ !->> _ /= _ _; rewrite pi_x2 oget_some /=. + rewrite domE; case: {-1}(ro0.[rcons p0 (v0 +^ x1)]) (eq_refl (ro0.[rcons p0 (v0 +^ x1)])). + + done. + move=> bo ^ro_pvx1 /=. have /mh_of_INV [] _ -> _:= inv0. + rewrite negb_exists=> ? /=; rewrite negb_exists=> ? /=; rewrite negb_exists=> yh /=. + rewrite Hpath /=; rewrite negb_and -implyNb /= => [#] !<<-. + rewrite xorwA xorwK xorwC xorw0 -negP=> G1mh_x1hx2. + have /m_mh_of_INV [] _ /(_ _ _ _ _ G1mh_x1hx2) := inv0. + move=> [xc xf yc yf] [#]; rewrite hs_hx2=> [#] <*>. + by rewrite PFm_x1x2. + auto => &m1 &m2 [#] !<- _ _ -> /= _ y1L ? y2L ? /=. + rewrite !get_set_sameE pi_x2 oget_some /=. + have /hs_of_INV [] Hu _ _:= inv0; have -> := huniq_hinvK_h _ _ _ Hu hs_hx2. + rewrite oget_some domE => /= ;have[]_->_ _ _/=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + case(G1.bcol{m2} \/ hinv hs0 y2L <> None)=>//=;rewrite !negb_or/==>[][]? hinv0[]? hinv1. + case:inv0=> Hhs Hinv HinvG Hmmh Hmmhi Hincl Hincli Hmh Hpi Hmp. + have Hhx2:= dom_hs_neq_ch _ _ _ _ _ Hhs hs_hx2. + have mh_hx2: G1mh.[(x1,hx2)] = None. + + case Hmmh => _ /(_ x1 hx2);case (G1mh.[(x1, hx2)]) => // -[ya hy] /(_ ya hy) /=. + by rewrite -negP=> -[xc fx yc fy];rewrite hs_hx2 => -[[!<<-]];rewrite PFm_x1x2. + have ch_0 := ch_neq0 _ _ Hhs. + have ch_None : forall xa xb ha hb, G1mh.[(xa,ha)] = Some(xb, hb) => ha <> ch0 /\ hb <> ch0. + + move=> xa xb ha hb;case Hmmh=> _ H /H [xc fx yc fy [#]]. + by move=> /(dom_hs_neq_ch _ _ _ _ _ Hhs) -> /(dom_hs_neq_ch _ _ _ _ _ Hhs). + split. + + by apply hs_addh => //;have /# := hinvP hs0 y2L. + + apply inv_addm=> //; case: {-1}(G1mi.[(y1L,y2L)]) (eq_refl G1mi.[(y1L,y2L)])=> //. + move=> [x1L x2L] ^G1mi_y; rewrite -Hincli 1:G1mi_y//. + case: Hmmhi hinv0 => H _ + /H {H} [hx fx hy fy] [#]. + by case: (hinvP hs0 y2L)=> [_ ->|//]/#. + + by apply inv_addm=>//; apply (ch_notdomE2_mh _ _ Hmmhi Hhs). + + by apply (m_mh_addh_addm _ Hmmh _ hs_hx2)=>//;apply ch_notdomE_hs. + + apply (mi_mhi_addh_addmi _ Hmmhi _ hs_hx2);last by apply ch_notdomE_hs. + by have := hinvP hs0 y2L;rewrite /#. + + by apply incl_addm. + by apply incl_addm. + + split. + + move=> xa hx ya hy;rewrite get_setE;case ((xa, hx) = (x1, hx2))=> /=. + + move=> [] !->> [] !<<-; exists x2 Known y2L Known. + by rewrite /= !get_set_sameE /= get_set_neqE. + move=> Hdiff Hxa; case Hmh=> /(_ _ _ _ _ Hxa) [] xc fx yc fy [#] Hhx Hhy HG1 _ _. + exists xc fx yc fy;rewrite !get_set_neqE //. + + by apply (dom_hs_neq_ch _ _ _ Hhs Hhx). + + by apply (dom_hs_neq_ch _ _ _ Hhs Hhy). + + rewrite /= -negP=> -[] <<- <<-;apply Hdiff=> /=. + by apply (Hu hx (xc, fx) (xc, Known)). + rewrite Hhx Hhy=> /=;move: HG1. + case: fy Hhy=> Hhy //= [p v [Hro Hbu]]. + exists p v;split. + + rewrite get_set_neqE // -negP => ^ /rconssI <<- /rconsIs. + move: Hbu;rewrite Hpath /= => -[!<<-] /=. + by rewrite -negP=> /Block.WRing.addrI /#. + by apply build_hpath_up=> //; move: hs_hx2 PFm_x1x2;apply: m_mh_None. + + move=> p bn b; rewrite get_setE. + case (rcons p bn = rcons p0 (v0 +^ x1)). + + rewrite /= => ^ /rconssI <<- /rconsIs ->> /=; split => [<<- | ]. + + exists v0 hx2 ch0. + rewrite (build_hpath_up Hpath) /=;1:by move: hs_hx2 PFm_x1x2;apply: m_mh_None. + by rewrite xorwA xorwK Block.WRing.add0r get_set_sameE. + move=> [v hx hy] [];rewrite get_setE ;case ((v +^ (v0 +^ x1), hx) = (x1, hx2)) => //. + move=> Hdiff;have HG1 := m_mh_None _ _ _ _ _ _ _ Hmmh hs_hx2 PFm_x1x2. + have -> /= [->> <<-]:= build_hpath_up_None _ _ (y1L, ch0) _ _ HG1 Hpath. + by move:Hdiff;rewrite xorwA xorwK Block.WRing.add0r. + rewrite /= => Hdiff. + case Hmh => ? -> Huni. + apply exists_iff=> v /= ;apply exists_iff => hx /=;apply exists_iff => hy /=. + rewrite build_hpath_upd_ch_iff //. + case (hx = ch0) => [->>|?]. + + split;1: by move=> [] _ /ch_None. + move=> [[p0' x [Hhx2']]]. + have [!->>] [!->>]:= Huni _ _ _ _ _ Hpath Hhx2'. + rewrite get_set_neqE /=. + + by rewrite (@eq_sym ch0 hx2) Hhx2. + by move => /ch_None. + rewrite get_setE;case ((v +^ bn, hx) = (x1, hx2)) => //= -[<<- ->>]. + split=> -[H];have [!->>]:= Huni _ _ _ _ _ Hpath H;move:Hdiff; + by rewrite xorwA xorwK Block.WRing.add0r. + move=> p v p' v' hx;case Hmh => _ _ Huni. + rewrite !build_hpath_upd_ch_iff //. + case (hx = ch0) => [->> [?? [# H1 -> ->]] [?? [# H2 -> ->]]|_ ] /=. + + by have [!->>] := Huni _ _ _ _ _ H1 H2. + by apply Huni. + split=> c p v;rewrite get_setE. case (c = y2L) => [->> /= | Hc]. + + split. + + move=> [!<<-];exists ch0;rewrite get_set_sameE /= build_hpath_prefix. + exists v0 hx2;rewrite xorwA xorwK Block.WRing.add0r get_set_sameE /=. + have HG1 := m_mh_None _ _ _ _ _ _ _ Hmmh hs_hx2 PFm_x1x2. + by apply build_hpath_up_None. + move=> [h []];rewrite get_setE build_hpath_upd_ch_iff //. + case (h=ch0)=> [->> /= [??[# H1 -> ->]]| Hh] /=. + + by case Hmh => _ _ /(_ _ _ _ _ _ Hpath H1). + by have := hinvP hs0 y2L;rewrite /= => /#. + case Hpi => ->;apply exists_iff => h /=. + rewrite build_hpath_upd_ch_iff // get_setE;case (h = ch0) => [->> | //]. + split;1: by move=> [_ /(dom_hs_neq_ch _ _ _ _ _ Hhs)]. + by move=> /= [_ <<-];move:Hc. + split. + + by have[]/#:=Hmp. + + by have[]/#:=Hmp. + + have[]_ _ h _ _ l hdom i hi:=Hmp. + have[]b c[]->h':=h l hdom i hi. + by exists b c=>//=;rewrite get_setE/=-h';smt(domE take_oversize). + + by have[]/#:=Hmp. + + by have[]/#:=Hmp. + move=> [xa xc] PFm_x1x2. rcondf{1} 1; 1:by auto=> &hr [#] !<<- _ _ ->>; rewrite domE PFm_x1x2. + have /m_mh_of_INV [] + _ - /(_ _ _ _ _ PFm_x1x2) := inv0. + move=> [hx2 fx2 hy2 fy2] [#] hs_hx2 hs_hy2 G1mh_x1hx2. + case @[ambient]: {-1}(G1m.[(x1,x2)]) (eq_refl (G1m.[(x1,x2)])); last first. + + move=> [ya yc] G1m_x1x2; rcondf{2} 1; 1:by auto=> &hr [#] !<<- _ _ ->>; rewrite domE G1m_x1x2. + auto=> &1 &2 [#] <*> -> -> -> /=; have /incl_of_INV /(_ (x1,x2)) := inv0. + by rewrite PFm_x1x2 G1m_x1x2. + move=> x1x2_notin_G1m; rcondt{2} 1; 1:by auto=> &hr [#] !<<- _ _ ->>; rewrite domE x1x2_notin_G1m. + have <*>: fy2 = Unknown. + + have /mh_of_INV [] /(_ _ _ _ _ G1mh_x1hx2) + _ := inv0. + move=> [xc0 xf0 yc0 yf0] [#]; rewrite hs_hx2 hs_hy2=> [#] !<<- [#] !<<-. + by case: fy2 hs_hy2 G1mh_x1hx2=> //=; rewrite x1x2_notin_G1m. + case @[ambient]: fx2 hs_hx2=> hs_hx2. + + swap{2} 3 -2; seq 0 1: (queries = C.queries{2} /\ G1.bext{2}). + - by auto=> ? ? [#] !<<- _ -> ->> _ /=; rewrite rngE; exists hx2. + conseq(:_==> (! (G1.bcol{2} \/ G1.bext{2})) => oget PF.m{1}.[x{1}] = y{2} /\ + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} C.queries{2}); + progress;2..-2:rewrite/#. + - by rewrite domE;have[]_->_ _ _/=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + by inline*; if{2}; auto=> &1 &2 />; smt(F.sampleto_ll sampleto_ll). + have /mh_of_INV []/(_ _ _ _ _ G1mh_x1hx2) + _ _:= inv0. + move=> [xc0 xf0 yc0 yf0] [#]; rewrite hs_hx2 hs_hy2=> [#] !<<- [#] !<<- /= [p0 v0] [#] Hro Hpath. + have /pi_of_INV [] /(_ x2 p0 v0) /iffRL /(_ _) := inv0. + + by exists hx2=>/#. + move=> pi_x2; rcondt{2} 1; 1:by auto=> &hr [#] <*>; rewrite domE pi_x2. + inline F.RO.get. + rcondf{2} 4; first by auto=> &hr [#] !<<- _ _ ->> _ /=; rewrite pi_x2 oget_some /= domE Hro. + rcondf{2} 8; first by auto=> &hr [#] !<<- _ _ ->> _ /= _ _ _ _; rewrite rngE; exists hx2. + rcondt{2} 9. + + auto=> &hr [#] !<<- _ _ ->> _ /= _ _ _ _. + rewrite (@huniq_hinvK_h hx2 hs0 x2) // 2:domE 2:G1mh_x1hx2 2:!oget_some /=. + + by have /hs_of_INV []:= inv0. + by rewrite /in_dom_with domE hs_hy2. + rcondt{2} 14; first by auto=> &hr [#] !<<- _ _ ->> _ /=; rewrite domE pi_x2. + auto=> &1 &2 [#] !<<- -> -> ->> _ /=. + move=> _ _ _ _; rewrite PFm_x1x2 pi_x2 !oget_some //=. + rewrite (@huniq_hinvK_h hx2 hs0 x2) // 10?oget_some. + + by have /hs_of_INV []:= inv0. + rewrite Hro G1mh_x1hx2 hs_hy2 ?oget_some //=domE. + have[]_->_ _ _//=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ inv0. + case(rng hs0 (x2, Unknown))=>//=_. + exact/(@lemma3 _ _ _ _ _ _ _ _ _ _ _ _ _ _ hx2 _ _ hy2). + by move=> /> &1 &2 -> ->. +qed. + +lemma head_nth (w:'a) l : head w l = nth w l 0. +proof. by case l. qed. + +lemma drop_add (n1 n2:int) (l:'a list) : 0 <= n1 => 0 <= n2 => drop (n1 + n2) l = drop n2 (drop n1 l). +proof. + move=> Hn1 Hn2;elim: n1 Hn1 l => /= [ | n1 Hn1 Hrec] l;1: by rewrite drop0. + by case: l => //= a l /#. +qed. + +lemma behead_drop (l:'a list) : behead l = drop 1 l. +proof. by case l => //= l;rewrite drop0. qed. + +lemma incl_upd_nin (m1 m2:('a,'b)fmap) x y: incl m1 m2 => x \notin m2 => incl m1 m2.[x <- y]. +proof. + move=> Hincl Hdom w ^/Hincl <- => Hw. + rewrite get_set_neqE // -negP => ->>. + by move: Hdom;rewrite domE. +qed. + + + +lemma lemma5 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries i (p : block list) b c h: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries + => 0 <= i < size p + => take (i + 1) p \in prefixes + => prefixes.[take i p] = Some (b,c) + => (exists f, hs.[h] = Some (c,f)) + => exists b' c' h', + Pm.[(b +^ nth witness p i, c)] = Some (b',c') /\ + mh.[(b +^ nth witness p i, h)] = Some (b',h'). +proof. +move=>Hinv H_size H_take_iS H_take_i H_hs_h. +have[]_ _ H _ _:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. +have[]sa sc:=H _ H_take_iS i _;1:smt(size_take). +rewrite!take_take !minrE (: i <= i + 1) 1: /# nth_take 1,2:/#H_take_i=>[][]/=[->>->>] H_pm. +have[]b' c' H_Pm:exists b' c', Pm.[(sa +^ nth witness p i, sc)] = Some (b',c'). ++ rewrite H_pm. exists (oget prefixes.[take (i + 1) p]).`1 (oget prefixes.[take (i + 1) p]).`2. + by move: H_take_iS; rewrite domE; case: (prefixes.[take (i + 1) p])=> //= - []. +exists b' c';rewrite -H_Pm/=. +have[]h_Pm _:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. +have[]h' f' hy fy[]H_h'[]H_hy H_mh:=h_Pm _ _ _ _ H_Pm. +have[]h_huniq _ _:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. +have[]f H_h := H_hs_h. +have/=<<-:=h_huniq _ _ _ _ H_h H_h'. +by rewrite H_mh/=/#. +qed. + + +lemma lemma5' hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries i (p : block list) b c h: + INV_CF_G1 hs ch Pm Pmi Gm Gmi mh mhi ro pi prefixes queries + => 0 <= i < size p + => prefixes.[take i p] = Some (b,c) + => (exists f, hs.[h] = Some (c,f)) + => (exists b' c' h', + Pm.[(b +^ nth witness p i, c)] = Some (b',c') /\ + mh.[(b +^ nth witness p i, h)] = Some (b',h')) \/ + (Pm.[(b +^ nth witness p i, c)] = None /\ + mh.[(b +^ nth witness p i, h)] = None). +proof. +move=>Hinv H_size H_take_i H_hs_h. +case(Pm.[(b +^ nth witness p i, c)] = None)=>//=H_Pm. ++ right;move:H_Pm;apply absurd=>H_mh. + have[]b1 h1 H_mh1:exists b1 h1, mh.[(b +^ nth witness p i, h)] = Some (b1,h1). + + exists (oget mh.[(b +^ nth witness p i, h)]).`1 (oget mh.[(b +^ nth witness p i, h)]).`2. + by move: H_mh; case: (mh.[(b +^ nth witness p i, h)])=> //= - []. + have[]H_Pm H_Gmh:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. + by have/#:=H_Gmh _ _ _ _ H_mh1. +have[]b1 c1 H_Pm1:exists b1 c1, Pm.[(b +^ nth witness p i, c)] = Some (b1,c1) + by exists (oget Pm.[(b +^ nth witness p i, c)]).`1 + (oget Pm.[(b +^ nth witness p i, c)]).`2;smt(domE). +have[]H_P_m H_Gmh:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. +have:=H_P_m _ _ _ _ H_Pm1. +have[] :=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ Hinv. +move=> hun *. +have /> := H_P_m _ _ _ _ H_Pm1. +move=> hx fx hy fy H1 H2 H3; exists b1 c1 hy => />. +case: H_hs_h => fh /(hun _ _ _ _ H1) />. +qed. + + +equiv PFf_Cf (D<:DISTINGUISHER): + + DFRestr(SqueezelessSponge(PF)).f ~ DFRestr(G1(DRestr(D)).M).f : + + ! (G1.bcol{2} \/ G1.bext{2}) /\ + ={arg} /\ ={glob C} /\ + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} C.queries{2} + ==> + if G1.bcol{2} \/ G1.bext{2} + then ([] \in C.queries{2}) + else ={glob C} /\ ={res} /\ + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} C.queries{2}. +proof. + proc;sp;inline*;sp. + if;1,3:auto;if;1,3:auto;swap{1}4;swap{2}11;sp;wp 1 5. + sp;conseq(:_==> ! (G1.bcol{2} \/ G1.bext{2}) => + ={glob C, sa} /\ + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} G1.mh{2} + G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} C.queries{2}.[bs{1} <- sa{1}] /\ + F.RO.m.[p]{2} = Some sa{2});progress. + + by rewrite mem_set domE;left;have[]_->_ _ _//=:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0. + + smt(mem_set). + + smt(mem_set). + + smt(mem_set). + seq 1 1: + (={i, p, glob C} /\ i{1} = size p{1} /\ p{2} = bs{1} /\ + (!(G1.bcol{2} \/ G1.bext{2}) => + (INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2}.[bs{1} <- sa{1}] + /\ ={sa} /\ F.RO.m.[p]{2} = Some sa{1})));last first. + + case : (! (G1.bcol{2} \/ G1.bext{2}));last first. + - by conseq(:_==>true);progress;auto;smt(DBlock.dunifin_ll DCapacity.dunifin_ll take_size). + by rcondf{2}3;auto;smt(domE DBlock.dunifin_ll DCapacity.dunifin_ll take_size). + + conseq(:_==> ={i, p, glob C} /\ i{1} = size p{1} /\ p{2} = bs{1} /\ + (!(G1.bcol{2} \/ G1.bext{2}) => + (INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2}.[take i{2} bs{1} <- sa{1}] + /\ ={sa} /\ F.RO.m.[p]{2} = Some sa{1})));1:smt(take_size). + + splitwhile{1} 1 : i < prefix p (get_max_prefix p (elems (fdom C.queries))). + splitwhile{2} 1 : i < prefix p (get_max_prefix p (elems (fdom C.queries))). + + seq 1 1 : (={p, i, glob C, bs} /\ bs{2} = p{2} /\ + (prefix p (get_max_prefix p (elems (fdom C.queries))) = i){2} /\ + (Redo.prefixes.[take i p]{1} = Some (sa,sc){1}) /\ + (take i p \in Redo.prefixes){1} /\ + (C.queries.[[]] = Some b0){1} /\ + (! p{2} \in C.queries{2}) /\ + (!(G1.bcol{2} \/ G1.bext{2}) => + (INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2} /\ + ={sa} /\ counter{2} = 0 /\ + (exists f, FRO.m.[h]{2} = Some (sc{1}, f)) /\ + (build_hpath G1.mh (take i p) = Some (sa,h)){2} /\ + if i{2} = 0 then (sa,h){2} = (b0, 0) + else F.RO.m.[take i p]{2} = Some sa{1})) /\ + (i{2} = 0 => sa{1} = b0) /\ 0 < size p{2}). + + while(={p, i, glob C} /\ bs{2} = p{2} /\ (i{2} = 0 => sa{1} = b0) /\ + (0 <= i <= prefix p (get_max_prefix p (elems (fdom C.queries)))){2} /\ + (Redo.prefixes.[take i p]{1} = Some (sa,sc){1}) /\ + (take i p \in Redo.prefixes){1} /\ + (C.queries.[[]] = Some b0){1} /\ + (! p{2} \in C.queries{2}) /\ + (!(G1.bcol{2} \/ G1.bext{2}) => + (INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2} /\ + ={sa} /\ counter{2} = 0 /\ + (exists f, FRO.m.[h]{2} = Some (sc{1}, f)) /\ + (build_hpath G1.mh (take i p) = Some (sa,h)){2} /\ + if i{2} = 0 then (sa,h){2} = (b0, 0) + else F.RO.m.[take i p]{2} = Some sa{1})) /\ 0 < size p{2});last first. + - auto;progress. + * smt(@Prefix). + * by have[]:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0;smt(take0 domE). + * by have[]:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0;smt(take0 domE). + * by have[]:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0. + * by have[]:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0;smt(take0 domE). + * by rewrite build_hpathP; apply/Empty=> //; exact/take0. + * by have[]:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0; smt(take0 domE size_take size_eq0 size_ge0). + * smt(prefix_sizel). + + case(G1.bcol{2} \/ G1.bext{2}). + - by if{1};auto;conseq(:_==> (G1.bcol{2} \/ G1.bext{2}));1,3:smt(domE get_setE); + (if{2};2:if{2});auto;1:smt(DBlock.dunifin_ll DCapacity.dunifin_ll); + sp;if{1};auto;smt(DBlock.dunifin_ll DCapacity.dunifin_ll). + conseq(: ={p, i, glob C} /\ bs{2} = p{2} /\ (i{2} = 0 => sa{1} = b0) /\ + 0 <= i{2} <= prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) /\ + Redo.prefixes{1}.[take i{1} p{1}] = Some (sa{1}, sc{1}) /\ + (C.queries.[[]] = Some b0){1} /\ (! p{2} \in C.queries{2}) /\ + (take i{1} p{1} \in Redo.prefixes{1}) /\ + (! (G1.bcol{2} \/ G1.bext{2}) => + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2} /\ + ={sa} /\ + counter{2} = 0 /\ + (exists (f : flag), FRO.m{2}.[h{2}] = Some (sc{1}, f)) /\ + build_hpath G1.mh{2} (take i{2} p{2}) = Some (sa{2}, h{2}) /\ + if i{2} = 0 then (sa{2}, h{2}) = (b0, 0) + else F.RO.m{2}.[take i{2} p{2}] = Some sa{1}) /\ + (i{1} < size p{1} /\ + i{1} < prefix p{1} (get_max_prefix p{1} (elems (fdom C.queries{1})))) /\ + i{2} < size p{2} /\ + i{2} < prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) /\ + ! (G1.bcol{2} \/ G1.bext{2}) /\ (take (i+1) p \in Redo.prefixes){1} /\ + 0 < size p{2} + ==>_);progress. + - have:=prefix_gt0_mem p{2} (elems (fdom C.queries{2})) _;1:rewrite/#. + rewrite-memE=>H_dom_q. + have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H12. + have[]_ _ h1 h2:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have:=h2 (get_max_prefix p{2} (elems (fdom C.queries{2}))) _; 1:smt(mem_fdom). + move=>[]c; + have H_dom_p:get_max_prefix p{2} (elems (fdom C.queries{2})) \in Redo.prefixes{1} by smt(domE mem_fdom). + have->/=:=prefix_take_leq p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) (i{2}+1) _;1:rewrite/#. + smt(domE take_oversize prefix_sizer). + rcondt{1}1;1:auto;progress. + rcondt{2}1;1:auto;progress. + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H11. + have//=:=lemma5 _ _ _ _ _ _ _ _ _ _ _ _ i{hr} p{hr} sa{hr} sc{m} h{hr} HINV _ _ _ _. + * by rewrite H0/=H7/=. + * smt(domE). + * rewrite/#. + * rewrite/#. + by move=>[]b2 c2 h2[]H_PFm H_Gmh;rewrite domE H_Gmh/=. + auto;progress. + - rewrite /#. + - rewrite /#. + - rewrite /#. + - smt(domE). + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO/#:=H6 H11. + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H11. + have[]H01 H02 H_pref1 H_pref2:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have//=:=lemma5 _ _ _ _ _ _ _ _ _ _ _ _ i{2} p{2} sa{2} sc{1} h{2} HINV _ _ _ _. + * by rewrite H0/=H7/=. + * smt(domE). + * rewrite/#. + * rewrite/#. + move=>[]b2 c2 h2[]H_PFm H_Gmh. + rewrite H_Gmh/=. + have[]b6 c6[]:=H_pref1 _ H12 i{2} _;1:smt(size_take). + by rewrite!take_take !minrE (: i{2} <= i{2} + 1) //= 1:/# nth_take 1,2:/# H2/==>[][]->>->><-;rewrite H_PFm oget_some. + - rewrite/#. + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H11. + have[]H01 H02 H_pref1 H_pref2:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have//=:=lemma5 _ _ _ _ _ _ _ _ _ _ _ _ i{2} p{2} sa{2} sc{1} h{2} HINV _ _ _ _. + * by rewrite H0/=H7/=. + * smt(domE). + * rewrite/#. + * rewrite/#. + move=>[]b2 c2 h2[]H_PFm H_Gmh. + have[]b6 c6[]:=H_pref1 _ H12 i{2} _;1:smt(size_take). + rewrite!take_take !minrE (: i{2} <= i{2} + 1) // 1:/# nth_take 1,2:/# H2/=H_Gmh oget_some=>[][]<<-<<-<-. + rewrite H_PFm oget_some/=. + have [] help1 help2:= m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + by have [] xc fx yc fy [#] /# := help2 _ _ _ _ H_Gmh. + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H11. + have[]H01 H02 H_pref1 H_pref2:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have//=:=lemma5 _ _ _ _ _ _ _ _ _ _ _ _ i{2} p{2} sa{2} sc{1} h{2} HINV _ _ _ _. + * by rewrite H0/=H7/=. + * smt(domE). + * rewrite/#. + * rewrite/#. + move=>[]b2 c2 h2[]H_PFm H_Gmh. + by rewrite H_Gmh/= (@take_nth witness) 1:/# build_hpath_prefix/#. + - rewrite/#. + - rewrite/#. + - have[]HINV[]->>/=[]->>/=[]H_h[]H_path H_F_RO:=H6 H11. + have[]H01 H02 H_pref1 H_pref2:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have//=:=lemma5 _ _ _ _ _ _ _ _ _ _ _ _ i{2} p{2} sa{2} sc{1} h{2} HINV _ _ _ _. + * by rewrite H0/=H7/=. + * smt(domE). + * rewrite/#. + * rewrite/#. + move=>[]b2 c2 h2[]H_PFm H_Gmh. + have[]b6 c6[]:=H_pref1 _ H12 i{2} _;1:smt(size_take). + rewrite!take_take !minrE (: i{2} <= i{2} + 1) 1:/# nth_take 1,2:/# H2/==>[][]<<-<<-<-. + rewrite H_PFm/=(@take_nth witness)1:/#. + by have[]help1 help2/# :=mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + alias{1} 1 prefixes = Redo.prefixes;sp. + alias{2} 1 bad1 = G1.bcol;sp. + + while ( ={i, p, C.queries, C.c} + /\ prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) <= + i{1} <= size p{1} + /\ Redo.prefixes{1}.[take i{2} p{2}] = Some (sa{1}, sc{1}) + /\ p{2} = bs{1} + /\ (! p{2} \in C.queries{2}) + /\ (! (G1.bcol{2} \/ G1.bext{2}) => + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} G1.m{2} G1.mi{2} + G1.mh{2} G1.mhi{2} F.RO.m{2} G1.paths{2} Redo.prefixes{1} + C.queries{2}.[take i{2} bs{1} <- sa{1}] + /\ ! (bad1{2} \/ G1.bext{2}) + /\ m_p PF.m{1} prefixes{1} C.queries{2} + /\ (forall (l : block list), + l \in prefixes{1} => prefixes{1}.[l] = Redo.prefixes{1}.[l]) + /\ (forall (l : block list), l \in Redo.prefixes{1} => + (l \in prefixes{1}) \/ + exists (j : int), 0 <= j <= i{2} /\ take j p{2} = l) + /\ ={sa} + /\ counter{2} <= i{2} - prefix p{2} + (get_max_prefix p{2} (elems (fdom C.queries{2}))) + /\ (exists (f : flag), FRO.m{2}.[h{2}] = Some (sc{1}, f)) + /\ build_hpath G1.mh{2} (take i{2} p{2}) = Some (sa{2}, h{2}) + /\ (if i{2} = 0 then (sa{2}, h{2}) = (b0, 0) + else F.RO.m{2}.[take i{2} p{2}] = Some sa{1}) + /\ (i{2} < size p{2} => ! take (i{2}+1) p{2} \in Redo.prefixes{1})));last first. + + auto;progress. + - smt(prefix_sizel). + - have[]HINV [#] ->> _ _ _ h_sa_b0:=H3 H6;split;..-2:case:HINV=>//=. + have[]Hmp01 Hmp02 Hmp1 Hmp2 Hmp3:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV; split=> //=. + + move: h_sa_b0; case: (prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2}))) = 0). + + by move=> -> [#] ->> _; rewrite take0 get_set_sameE. + smt(size_take get_setE). + + move=> l; rewrite mem_set=> - []. + + move=> /Hmp2 [c] h. + case: (l = take (prefix p{2} (get_max_prefix p{2} (elems (fdom C.queries{2})))) p{2}). + + by move=> ->>; exists sc{1}; rewrite get_set_sameE H. + by move=> n_not_crap; exists c; rewrite get_set_neqE. + by move=> ->>; exists sc{1}; rewrite get_set_sameE H. + by move=> l /Hmp3 [l2] ll2_in_q; exists l2; rewrite mem_set ll2_in_q. + - by have[]HINV _:=H3 H6;have:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - have[]HINV[]->>[]->>[]H_h[]H_path H_F_RO:=H3 H6. + have[]H01 H02 Hmp1 Hmp2 Hmp3:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_pref_eq:=prefix_exchange_prefix_inv (elems (fdom C.queries{2})) + (elems (fdom Redo.prefixes{1})) p{2} _ _ _. + * smt(memE domE mem_fdom). + * smt(memE mem_fdom domE take_oversize size_take take_take nth_take take_le0). + * smt(memE mem_fdom domE take_oversize size_take take_take nth_take take_le0). + by rewrite -mem_fdom memE prefix_lt_size 1:-H_pref_eq /#. + - rewrite/#. + - rewrite/#. + - rewrite/#. + - smt(take_size). + + case : (! (G1.bcol{2} \/ G1.bext{2}));last first. + - wp 1 1=>/=. + conseq(:_==> Redo.prefixes{1}.[take (i{1}+1) p{1}] = Some (sa{1}, sc{1}) + /\ (take (i{1} + 1) p{1} \in Redo.prefixes{1}) + /\ (G1.bcol{2} \/ G1.bext{2}));1:smt(prefix_ge0). + if {1}; last first. + + sp; if {1}; if {2}; last first. + + if {2}; sp; auto=> />. + + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + + if {2}; last first. + + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + + swap {2} 4 -3; auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + if {2}. + + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + if {2}. + + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + auto=> />. + smt(get_setE mem_set DBlock.dunifin_ll DCapacity.dunifin_ll). + + rcondf{1}1;1:auto=>/#. + sp;wp. + if{1};2:rcondt{2}1;first last;3:rcondf{2}1;..3:auto. + + progress. + have[]HINV[]Hbad[]HINV0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{hr} bs{m} sa{hr} sc{m} h{hr} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + by have:=H7;rewrite !domE=>->/=/#. + + progress. + - rewrite/#. + - rewrite/#. + - by rewrite get_setE. + - have[]HINV[]Hbad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + split;..-2:case:HINV=>//=. + have[]Hmp01 Hmp02 Hmp1 Hmp2 Hmp3:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV;split=>//=. + * smt(get_setE size_take size_eq0 size_ge0 prefix_ge0). + * by have[]_ Hmp02' _ _ _:=H_m_p0; + smt(get_setE size_take size_eq0 size_ge0 prefix_ge0 take0). + * move=>l;rewrite!mem_set. + case(l = take (i{2} + 1) bs{1})=>//=[->>|]. + + move=>j;rewrite size_take;1:smt(prefix_ge0). + have->/=:(if i{2} + 1 < size bs{1} then i{2} + 1 else size bs{1}) = i{2} + 1 by rewrite/#. + move=>[]H0j HjiS;rewrite!get_setE. + have->/=:! take j (take (i{2} + 1) bs{1}) = take (i{2} + 1) bs{1} by smt(size_take). + rewrite!take_take!minrE (: j <= i{2} + 1) 1:/# (: j + 1 <= i{2} + 1) 1:/#. + rewrite nth_take 2:/#;1:smt(prefix_ge0). + case(j < i{2})=>Hij. + - have->/=:!take (j + 1) bs{1} = take (i{2} + 1) bs{1} by smt(size_take). + by have:=Hmp1(take i{2} bs{1}) _ j _; + smt(domE take_take nth_take prefix_ge0 size_take). + have->>:j = i{2} by rewrite/#. + by exists sa{2} sc{1};rewrite H1/=;smt(). + move=>h H_dom j []Hi0 Hisize;rewrite!get_setE. + have->/=:!take j l = take (i{2} + 1) bs{1} by smt(domE take_oversize size_take take_take). + by have->/=/#:!take (j+1) l = take (i{2} + 1) bs{1} + by smt(domE take_oversize size_take take_take). + * move=>l;rewrite mem_set. + case(l = take (i{2} + 1) bs{1})=>//=[->>|]. + + by rewrite!get_setE/= /#. + move=>h H_dom;rewrite!get_setE h/=. + have[]H2mp01 H2mp02 H2mp1 H2mp2 H2mp3:=H_m_p0. + rewrite-Hp1;1:smt(domE). + by apply H2mp2. + move=>l;rewrite !mem_set. + case(l = take (i{2} + 1) bs{1})=>//=[->>|]. + + by exists []; smt(cats0 mem_set). + move=>H_neq H_dom;have[]l1:=Hmp3 _ H_dom;rewrite!mem_set;case=>H_case. + + exists l1;by rewrite mem_set H_case. + exists (rcons l1 (nth witness bs{1} i{2}));rewrite mem_set;right. + by rewrite-rcons_cat (@take_nth witness);smt(prefix_ge0). + - rewrite/#. + - rewrite/#. + - smt(domE get_setE). + - move:H9;rewrite mem_set;case;smt(prefix_ge0). + - have[]HINV[]Hbad[]HINV0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + by have:=H7;rewrite !domE=>->/=/#. + - rewrite/#. + - have[]HINV[]Hbad[]HINV0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + have:=H7;rewrite !domE=>->/=[]b4 c4 h4[]H_PFm H_Gmh;rewrite H_PFm H_Gmh !oget_some/=. + have[]_ help:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have:=help _ _ _ _ H_Gmh. + by have[]f H_h':=H_h;rewrite H_h'/==>[][]a b c d[][]->>->>[];rewrite H_PFm/==>[]h'->>/#. + - have[]HINV[]Hbad[]HINV0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + have:=H7;rewrite !domE=>->/=[]b4 c4 h4[]H_PFm H_Gmh. + rewrite (@take_nth witness);1:smt(prefix_ge0). + by rewrite build_hpath_prefix H_path/=;smt(domE). + - smt(prefix_ge0). + - smt(prefix_ge0). + - have[]HINV[]Hbad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + have:=H7;rewrite !domE=>->/=[]b4 c4 h4[]H_PFm H_Gmh. + rewrite(@take_nth witness);1:smt(prefix_ge0). + have[]_ help H_uniq_path:=mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + by rewrite help H_path;smt(domE). + - have[]HINV[]Hbad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + rewrite mem_set negb_or/=;split;2:smt(size_take prefix_ge0 take_oversize). + have:=Hp2 (take (i{2} + 1 + 1) bs{1}). + pose P:= _ \/ _;have/#:!P;rewrite/P;clear P;rewrite negb_or/=negb_exists/=;split. + * have:=prefix_exchange_prefix_inv(elems (fdom C.queries{2}))(elems (fdom prefixes{1}))bs{1} _ _ _. + + by have[]:=H_m_p0;smt(domE memE mem_fdom). + + have[]Hmp01 Hmp02 Hmp1 Hmp2 Hmp3:=H_m_p0. + have:=all_prefixes_of_m_p _ _ _ H_m_p0. + + move=> h_prefixes l2; rewrite -memE mem_fdom=> /Hmp2 [c]. + move=> pl2; move: (h_prefixes l2 _). + + by rewrite domE pl2. + by move=> + i - /(_ i); rewrite -memE mem_fdom. + + by have[]:=H_m_p0;smt(memE domE mem_fdom). + by move=>H_pref_eq;rewrite -mem_fdom memE prefix_lt_size//= -H_pref_eq/#. + by move=>j;case(0<=j<=i{2})=>//=[][]Hj0 Hji;smt(size_take prefix_ge0 take_le0). + + progress. + have[]HINV[]Hbad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_take_not_in:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{hr} bs{m} sa{hr} sc{m} h{hr} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - exact H_h. + by have:=H7;rewrite !domE=>/=->/=. + rcondt{2}1;1:auto=>/#. + rcondt{2}5;auto;progress. + * rewrite(@take_nth witness);1:smt(prefix_ge0);rewrite domE. + have[]HINV[]H_bad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[]H_h[]H_path[]H_F_RO H_i:=H3 H6. + have[]:=mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{hr} bs{m} sa{hr} sc{m} h{hr} HINV _ _ _. + * smt(prefix_ge0). + * rewrite/#. + * rewrite/#. + have:=H7;rewrite domE =>/=->/=H_Gmh _ H_ H_path_uniq. + have help:=H_ (take i{hr} bs{m}) (nth witness bs{m} i{hr});rewrite H_path/= in help. + have:forall (b : block), + F.RO.m{hr}.[rcons (take i{hr} bs{m}) (nth witness bs{m} i{hr})] = Some b + <=> exists hy, G1.mh{hr}.[(sa{hr} +^ nth witness bs{m} i{hr}, h{hr})] = Some (b, hy) by rewrite/#. + move:help=>_ help;move:H_Gmh;apply absurd=>//=H_F_Ro. + have[]b h: exists b, F.RO.m{hr}.[rcons (take i{hr} bs{m}) (nth witness bs{m} i{hr})] = Some b. + + by move: H_F_Ro; case: (F.RO.m{hr}.[rcons (take i{hr} bs{m}) (nth witness bs{m} i{hr})])=> //= /#. + by have:= (help b); rewrite h; smt(). + swap{2}-3;auto;progress. + * rewrite/#. + * rewrite/#. + * by rewrite!get_setE/=. + * have[]HINV[]H_bad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[][]f H_h[]H_path[]H_F_RO H_i:=H3 H6. + have:=H10;rewrite !negb_or/==>[][][]bad1 hinv_none bad2. + have H_hs_spec:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_mh_spec:=mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_m_mh:=m_mh_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_mi_mhi:=mi_mhi_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_pi_spec:=pi_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have :=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + * smt(prefix_ge0). + * exact H1. + * rewrite/#. + have:=H7;rewrite domE/==>->/=h_g1. + have H2_pi_spec:pi_spec FRO.m{2}.[G1.chandle{2} <- (y2L, Unknown)] + G1.mh{2}.[(sa{2} +^ nth witness bs{1} i{2}, h{2}) <- (y1L, G1.chandle{2})] + G1.paths{2}. + + split;progress. + - have[]h:=H_pi_spec;have:=h c p0 v;rewrite H11/==>[][]h1[] h'1 h'2. + exists h1;rewrite -h'2 get_setE/=. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec h'2. + by apply build_hpath_up=>//=. + move:H12;rewrite get_setE/==>hh0. + have h0_neq_ch:h0 <> G1.chandle{2} by rewrite/#. + have[]->:=H_pi_spec;rewrite-hh0 h0_neq_ch/=;exists h0=>/=. + have:=H;have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p0 v h0. + rewrite h_g1/=H/=h0_neq_ch/=. + have->//=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec H_h. + have -> /= <-//=:=ch_neq0 _ _ H_hs_spec;progress;have[]hh1 hh2 hh3:=H_mh_spec;smt(dom_hs_neq_ch). + split. + + apply hs_addh;1:have//:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + by have:=hinvP FRO.m{2} y2L;rewrite hinv_none/=/#. + + by have:=invG_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + apply inv_addm=>//;1:have//:=inv_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + apply (notin_hs_notdomE2_mh FRO.m{2} PF.mi{1})=>//=. + by apply ch_notdomE_hs;have:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + have[] H_huniq _ _:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + rewrite!get_setE/=. + apply (m_mh_addh_addm _ H_m_mh H_huniq H_h _)=>//=. + by apply ch_notdomE_hs;have:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + have[] H_huniq _ _:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + rewrite!get_setE/=;apply (mi_mhi_addh_addmi _ H_mi_mhi _ H_h _)=>//=. + - smt(hinvP). + by apply ch_notdomE_hs;have:=hs_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + apply incl_upd_nin=>//=. + by have:=incl_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + + apply incl_upd_nin=>//=. + - by have:=incli_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have:=hinvP FRO.m{2} y2L;rewrite domE hinv_none/=;apply absurd=>H_P_mi. + rewrite negb_forall/=. + have H_inv_Gmh:=inv_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have[]H_inv_Pm:=inv_mh_inv_Pm _ _ _ _ _ H_m_mh H_mi_mhi H_inv_Gmh. + have[]H_Pmi H_Gmhi:=mi_mhi_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + by have[]/#:=H_Pmi y1L y2L (oget PF.mi{1}.[(y1L, y2L)]).`1 + (oget PF.mi{1}.[(y1L, y2L)]).`2 _;1:smt(domE). + + have H_take_Si:=take_nth witness i{2} bs{1} _;1:smt(prefix_ge0). + split=>//=. + - move=>x hx y hy;rewrite !get_setE. + case((x, hx) = (sa{2} +^ nth witness bs{1} i{2}, h{2}))=>//=. + * move=>[->> ->>][<<- <<-]/=. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec H_h. + rewrite H_h/=. + exists sc{1} f y2L Unknown=>//=. + exists (take i{2} bs{1}) (sa{2})=>//=;rewrite get_setE Block.WRing.addKr/=. + rewrite/=(@take_nth witness)/=;1:smt(prefix_ge0). + by apply build_hpath_up=>//=;smt(domE). + move=> neq h1. + have[]hh1 hh2 hh3:=H_mh_spec. + have[]xc hxx yc hyc []h2[]h3 h4:=hh1 _ _ _ _ h1. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec h2. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec h3. + rewrite h2 h3/=;exists xc hxx yc hyc=>//=. + move:h4;case(hyc = Known)=>//=neq2[]p0 b[]hp0 hb. + exists p0 b;rewrite get_setE. + have->/=:=build_hpath_up _ _ _ y1L G1.chandle{2} _ _ _ hb h_g1. + have/#:!rcons p0 (b +^ x) = rcons (take i{2} bs{1}) (nth witness bs{1} i{2});move:neq;apply absurd=>//=h'. + have<<-:take i{2} bs{1}=p0 by rewrite/#. + have hbex:b +^ x = nth witness bs{1} i{2} by rewrite/#. + by have:=hb;rewrite H_path/==>[][->>->>]/=;rewrite-hbex Block.WRing.addKr/=. + - progress. + * move:H11;rewrite get_setE/=H_take_Si/=. + case(p0 = (take i{2} bs{1}))=>[->>|hpp0];rewrite!get_setE/=. + + have->/=:=build_hpath_up _ _ _ y1L G1.chandle{2} _ _ _ H_path h_g1. + case(bn = (nth witness bs{1} i{2}))=>[->> /= ->>|hbni]/=. + - by exists sa{2} h{2} G1.chandle{2}=>//=;rewrite get_setE/=. + have->/=:!rcons (take i{2} bs{1}) bn = rcons (take i{2} bs{1}) (nth witness bs{1} i{2}). + - move:hbni;apply absurd=>//=h. + exact/(rconsIs _ _ h). + move=>h_ro_p_bn. + have[]_ hh4 _:=H_mh_spec. + by have:=hh4 (take i{2} bs{1}) bn b0;rewrite h_ro_p_bn/=H_path/=;smt(get_setE @Block.WRing). + have->/=:!rcons p0 bn = rcons (take i{2} bs{1}) (nth witness bs{1} i{2}). + + move:hpp0;apply absurd=>/=h. + have:size p0 = size (take i{2} bs{1}) by smt(size_rcons). + move:h;pose p' := take i{2} bs{1};pose e := nth witness bs{1} i{2}. + by move=>h h';move:p0 p' h' bn e h;apply seq2_ind=>//=/#. + move=>h_ro_p_bn. + have[]_ hh4 _:=H_mh_spec. + have:=hh4 p0 bn b0;rewrite h_ro_p_bn/==>[][];progress. + have help:(sa{2} +^ nth witness bs{1} i{2}, h{2}) <> (v +^ bn, hx) by rewrite/#. + exists v hx hy=>//=;rewrite get_setE;rewrite eq_sym in help;rewrite help/=H12/=. + by apply build_hpath_up=>//=. + move:H11 H12;rewrite!get_setE/= =>h_build_hpath_set. + case(hy = G1.chandle{2})=>//=[->>|hy_neq_ch]/=. + + move=>h;have h_eq:v +^ bn = sa{2} +^ nth witness bs{1} i{2} && hx = h{2}. + + have/#:G1.mh{2}.[(v +^ bn, hx)] <> Some (b0, G1.chandle{2}). + have[]_ hh2:=H_m_mh. + have:=hh2 (v +^ bn) hx b0 G1.chandle{2}. + case(G1.mh{2}.[(v +^ bn, hx)] = Some (b0, G1.chandle{2}))=>//=. + rewrite negb_exists/=;progress; + rewrite negb_exists/=;progress; + rewrite negb_exists/=;progress; + rewrite negb_exists/=;progress;rewrite !negb_and. + by have[]/#:=H_hs_spec. + have[]eq_xor ->>:=h_eq. + move:h;rewrite eq_xor/==>->>. + have/#:!(p0 = (take i{2} bs{1}) /\ bn = (nth witness bs{1} i{2})) => + F.RO.m{2}.[rcons p0 bn] = Some b0. + move:H_h;case:f=>h_flag;last first. + - have:=known_path_uniq _ _ _ sc{1} h{2} p0 v (take i{2} bs{1}) sa{2} H2_pi_spec _ h_build_hpath_set _. + * rewrite get_setE/=h_flag. + by have->//=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec h_flag. + * by apply build_hpath_up=>//=. + move=>[]->>->>/=;apply absurd=>//=_. + have->:bn = sa{2} +^ sa{2} +^ bn;smt(@Block). + have[]hh1 hh2 hh3:=H_mh_spec. + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) b0 p0 v h{2}. + rewrite h_build_hpath_set/=h_g1/=. + have->/=:=ch_neq0 _ _ H_hs_spec. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec h_flag. + move=>help;have:= help _;1:smt(dom_hs_neq_ch). + move=>h_build_hpath_p0. + rewrite hh2 h_build_hpath_p0/==>h_neq. + exists v h{2}=>//=. + rewrite eq_xor h_g1/=;move:h_neq;apply absurd=>//=. + have:=hh3 _ _ _ _ _ H_path h_build_hpath_p0. + have->:bn = sa{2} +^ sa{2} +^ bn;smt(@Block). + move=>help;have h_neq:! (v +^ bn = sa{2} +^ nth witness bs{1} i{2} /\ hx = h{2}) by rewrite/#. + move: help. rewrite h_neq/==>h_g1_v_bn_hx. + have[]hh1 hh2 hh3:=H_mh_spec. + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p0 v hx. + rewrite h_build_hpath_set/=h_g1/=. + have->/=:=ch_neq0 _ _ H_hs_spec. + by have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec H_h; smt(dom_hs_neq_ch). + progress. + + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p0 v hx. + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p' v' hx. + move:H11 H12;rewrite!get_setE/= =>H13 H14;rewrite H13 H14. + have->/=:=ch_neq0 _ _ H_hs_spec. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec H_h. + rewrite h_g1/=. + have[]:=H_mh_spec => HH1 HH2 HH3 HH4 HH5. + have toto:(forall (xa xb : block) (ha hb : int), + G1.mh{2}.[(xa, ha)] = Some (xb, hb) => + ha <> G1.chandle{2} /\ hb <> G1.chandle{2}). + - move=> /> xa xb ha hb Hmhab. + have:=dom_hs_neq_ch FRO.m{2} G1.chandle{2} ha. + have:=dom_hs_neq_ch FRO.m{2} G1.chandle{2} hb. + have/#:FRO.m{2}.[ha] <> None /\ FRO.m{2}.[hb] <> None. + have[] _ HH6:=H_m_mh. + by have/> g1 g2 g3 g4 -> -> />:=HH6 _ _ _ _ Hmhab. + have{HH4}:=HH4 toto. + have{HH5}:=HH5 toto. + case: (hx = G1.chandle{2})=>[->>|hx_neq_ch/>]. + - move=>[] p1 x1 [#] hp11 ->> <<-. + move=>[] p2 x2 [#] hp21 ->> ->> /=. + pose y := (sa{2} +^ nth witness bs{1} i{2}). + pose y1:=x1+^y. pose y2:=x2+^y. + have[#]->>->>:=HH3 _ _ _ _ _ hp21 H_path. + by have[#]->><<-//=:=HH3 _ _ _ _ _ hp11 H_path. + move=>hp21 hp11. + by have[#]->>->>:=HH3 _ _ _ _ _ hp21 hp11. + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p0 v hx. + have:=build_hpath_upd_ch_iff h{2} G1.chandle{2} G1.mh{2} (sa{2} +^ nth witness bs{1} i{2}) y1L p' v' hx. + move:H11 H12;rewrite!get_setE/= =>H13 H14;rewrite H13 H14/=. + have->/=:=ch_neq0 _ _ H_hs_spec. + have->/=:=dom_hs_neq_ch _ _ _ _ _ H_hs_spec H_h. + rewrite h_g1/=. + have[]:=H_mh_spec => HH1 HH2 HH3 HH4 HH5. + have toto:(forall (xa xb : block) (ha hb : int), + G1.mh{2}.[(xa, ha)] = Some (xb, hb) => + ha <> G1.chandle{2} /\ hb <> G1.chandle{2}). + - move=> /> xa xb ha hb Hmhab. + have:=dom_hs_neq_ch FRO.m{2} G1.chandle{2} ha. + have:=dom_hs_neq_ch FRO.m{2} G1.chandle{2} hb. + have/#:FRO.m{2}.[ha] <> None /\ FRO.m{2}.[hb] <> None. + have[] _ HH6:=H_m_mh. + by have/> g1 g2 g3 g4 -> -> />:=HH6 _ _ _ _ Hmhab. + have{HH4}:=HH4 toto. + have{HH5}:=HH5 toto. + case: (hx = G1.chandle{2})=>[->>|hx_neq_ch/>]. + - move=>[] p1 x1 [#] hp11 ->> <<-. + move=>[] p2 x2 [#] hp21 ->> ->> /=. + pose y := (sa{2} +^ nth witness bs{1} i{2}). + pose y1:=x1+^y. pose y2:=x2+^y. + have[#]->>->>:=HH3 _ _ _ _ _ hp21 H_path. + by have[#]->><<-//=:=HH3 _ _ _ _ _ hp11 H_path. + move=>hp21 hp11. + by have[#]->>->>:=HH3 _ _ _ _ _ hp21 hp11. + + rewrite!get_setE/=;exact H2_pi_spec. + + rewrite!get_setE/=. + have H_m_p:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + have H_all_prefixes:=all_prefixes_of_INV _ _ _ _ _ _ _ _ _ _ _ _ HINV. + split;case:H_m_p=>//=Hmp01 Hmp02 Hmp1 Hmp2 Hmp3. + - smt(get_setE size_take prefix_ge0). + - by have[]:=H_m_p0;smt(get_setE size_take prefix_ge0). + - move=>l;rewrite mem_set;case=>H_case j []Hj0. + * move=>Hjsize;rewrite!get_setE/=. + have->/=:!take j l = take (i{2} + 1) bs{1} by rewrite/#. + have->/=:!take (j+1) l = take (i{2} + 1) bs{1} by rewrite/#. + smt(domE get_setE). + have->>:=H_case;rewrite size_take;1:smt(prefix_ge0). + have->/=:(if i{2} + 1 < size bs{1} then i{2} + 1 else size bs{1}) = i{2} + 1 by rewrite/#. + move=>HjiS;rewrite!get_setE. + have->/=:! take j (take (i{2} + 1) bs{1}) = take (i{2} + 1) bs{1} by smt(size_take). + rewrite!take_take!minrE (: j <= i{2} + 1) 1:/# (: j + 1 <= i{2} + 1) 1:/#. + rewrite nth_take 2:/#;1:smt(prefix_ge0). + case(j < i{2})=>Hij. + - have->/=:!take (j + 1) bs{1} = take (i{2} + 1) bs{1} by smt(size_take). + by have:=Hmp1(take i{2} bs{1}) _ j _;smt(domE take_take nth_take prefix_ge0 size_take get_setE). + have->>:j = i{2} by rewrite/#. + by exists sa{2} sc{1};rewrite H1/=;smt(get_setE domE). + - move=>l;rewrite mem_set. + case(l = take (i{2} + 1) bs{1})=>//=[->>|]. + + by rewrite!get_setE/= /#. + move=>h H_dom;rewrite!get_setE h/=. + have[]H2mp01 H2mp02 H2mp1 H2mp2 H2mp3:=H_m_p0. + rewrite-Hp1;1:smt(domE). + by apply H2mp2. + move=>l;rewrite !mem_set. + case(l = take (i{2} + 1) bs{1})=>//=[->>|]. + + by exists []; smt(cats0 mem_set). + move=>H_neq H_dom;have[]l1:=Hmp3 _ H_dom;rewrite!mem_set;case=>H_case. + + exists l1;by rewrite mem_set H_case. + exists (rcons l1 (nth witness bs{1} i{2}));rewrite mem_set;right. + by rewrite-rcons_cat (@take_nth witness);smt(prefix_ge0). + * rewrite/#. + * have[]HINV[]H_bad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[][]f H_h[]H_path[]H_F_RO H_i:=H3 H6. + split;have[]//= hmp01 hmp02 hmp1 hmp2 hmp3:=H_m_p0. + move=> l l_in_pref i hisize. + have//[] sa sc [#] pref_sasc pm_pref:= hmp1 l l_in_pref i hisize. + by exists sa sc; smt(get_setE domE take_take take_nth size_take + prefix_ge0 nth_take take_oversize take_le0 mem_fdom fdom_set). + + rewrite!get_setE/=;smt(domE). + + smt(get_setE domE take_take size_take prefix_ge0 nth_take take_oversize take_le0). + + rewrite!get_setE/=;smt(domE). + + rewrite/#. + + by rewrite!get_setE/=/#. + + rewrite!get_setE/=(@take_nth witness);1:smt(prefix_ge0);rewrite build_hpath_prefix. + have[]HINV[]H_bad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[][]f H_h[]H_path[]H_F_RO H_i:=H3 H6. + have:=lemma5' _ _ _ _ _ _ _ _ _ _ _ _ i{2} bs{1} sa{2} sc{1} h{2} HINV _ _ _. + - smt(prefix_ge0). + - exact H1. + - rewrite/#. + have:=H7;rewrite domE=>/=->/=H_Gmh. + have->/=:=build_hpath_up_None _ _ (y1L, G1.chandle{2})_ _ H_Gmh H_path;smt(get_setE). + + smt(prefix_ge0). + + smt(prefix_ge0). + + by rewrite!get_setE. + rewrite!mem_set negb_or/=;split;2:smt(prefix_ge0 size_take prefix_ge0 take_oversize). + have[]HINV[]H_bad[]H_m_p0[]Hp1[]Hp2[]->>[]H_counter[][]f H_h[]H_path[]H_F_RO H_i:=H3 H6. + have:=Hp2 (take (i{2} + 1 + 1) bs{1}). + pose P:= _ \/ _;have/#:!P;rewrite/P;clear P;rewrite negb_or/=negb_exists/=;split. + * have:=prefix_exchange_prefix_inv(elems (fdom C.queries{2}))(elems (fdom prefixes{1}))bs{1} _ _ _. + + by have[]:=H_m_p0;smt(domE memE mem_fdom). + + have[]Hmp01 Hmp02 Hmp1 Hmp2 Hmp3:=H_m_p0. + have:=all_prefixes_of_m_p _ _ _ H_m_p0. + move=> + l2; rewrite -memE mem_fdom=> + /Hmp2 [c] l2_in_q - /(_ l2 _). + + by rewrite domE l2_in_q. + by move=> + i - /(_ i); rewrite -memE mem_fdom. + + by have[]:=H_m_p0;smt(memE domE mem_fdom). + by move=>H_pref_eq;rewrite -mem_fdom memE prefix_lt_size//= -H_pref_eq/#. + by move=>j;case(0<=j<=i{2})=>//=[][]Hj0 Hji;smt(size_take prefix_ge0 take_le0). +qed. + + +section AUX. + + declare module D <: DISTINGUISHER {-PF, -RO, -G1, -Redo, -C}. + + declare axiom D_ll (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}): + islossless P.f => islossless P.fi => islossless F.f => + islossless D(F, P).distinguish. + + equiv CF_G1 : CF(DRestr(D)).main ~ G1(DRestr(D)).main: + ={glob D} ==> !(G1.bcol \/ G1.bext){2} => ={res}. + proof. + proc;inline*;wp. + call (_: G1.bcol \/ G1.bext, ={glob C} /\ + INV_CF_G1 FRO.m{2} G1.chandle{2} PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} G1.mh{2} G1.mhi{2} F.RO.m{2} + G1.paths{2} Redo.prefixes{1} C.queries{2}, + [] \in C.queries{2}). + (* lossless D *) + + exact/D_ll. + (** proofs for G1.S.f *) + (* equivalence up to bad of PF.f and G1.S.f *) + + conseq (_: !G1.bcol{2} + /\ !G1.bext{2} + /\ ={x, glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2} + ==> !G1.bcol{2} + => !G1.bext{2} + => ={res, glob C} + /\ INV_CF_G1 FRO.m{2} G1.chandle{2} + PF.m{1} PF.mi{1} + G1.m{2} G1.mi{2} G1.mh{2} G1.mhi{2} + F.RO.m{2} G1.paths{2} + Redo.prefixes{1} C.queries{2}). + + by move=> &1 &2; rewrite negb_or. + + progress;have[]:=m_p_of_INV _ _ _ _ _ _ _ _ _ _ _ _ H0;smt(domE). + (* For now, everything is completely directed by the syntax of + programs, so we can *try* to identify general principles of that + weird data structure and of its invariant. I'm not sure we'll ever + be able to do that, though. *) + conseq(eq_f D);progress=>/#. + (* lossless PF.f *) + + move=> &2 _; proc;inline*; sp;if=> //=; auto; sp;if;auto;smt (@Block.DBlock @Capacity.DCapacity). + (* lossless and do not reset bad G1.S.f *) + + move=> &1; proc; inline*;sp;if;auto;sp;if;auto. + conseq (_: _ ==> G1.bcol \/ G1.bext); 1:smt (). + inline *; if=> //=; wp; rnd predT; wp; rnd predT; auto. + + smt (@Block.DBlock @Capacity.DCapacity). + smt (@Block.DBlock @Capacity.DCapacity). + (** proofs for G1.S.fi *) + (* equiv PF.P.fi G1.S.fi *) + + conseq(eq_fi D)=>/#. + (* lossless PF.P.fi *) + + move=> &2 _; proc; inline*; sp; if; auto; sp; if; auto; smt (@Block.DBlock @Capacity.DCapacity). + (* lossless and do not reset bad G1.S.fi *) + + move=> &1; proc; inline*; sp; if; auto; sp; if;auto;smt (@Block.DBlock @Capacity.DCapacity). + (** proofs for G1.C.f *) + (* equiv PF.C.f G1.C.f *) + + conseq(PFf_Cf D);auto=>/#. + (* lossless PF.C.f *) + + move=> &2 _; proc; inline *; sp; if; auto; if; auto; while (true) (size p - i); auto. + + if; 1:auto=>/#. + sp; if; 2: auto=>/#. + by wp; do 2!rnd predT; auto; smt (size_behead @Block.DBlock @Capacity.DCapacity). + smt (size_ge0). + (* lossless and do not reset bad G1.C.f *) + + move=> _; proc; inline *; wp;sp;if;auto;sp;if;auto;sp. + conseq(:_==> true) (: _ ==> G1.bcol \/ G1.bext)=> //=. + + by move=> />. + + smt(@DBlock @DCapacity mem_set). + + while (G1.bcol \/ G1.bext)=> //=. + if; 1:by auto. + if;2:by auto. + by auto=> /> &hr [->|->]. + while (true) (size p - i)=> [z|]. + + if; 1:by auto=> /#. + if; 2:by auto=> /#. + by wp; rnd predT; wp; rnd predT; auto=> />; smt(@DBlock @DCapacity). + by auto=> /#. + (* Init ok *) + inline *; auto=> />; split=> [|/#]. + do !split. + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + + smt (get_setE emptyE build_hpath_map0). + smt (get_setE emptyE build_hpath_map0). + qed. + + +end section AUX. + +section. + + declare module D <: DISTINGUISHER{-Perm, -C, -PF, -G1, -RO, -Redo}. + + declare axiom D_ll (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}): + islossless P.f => islossless P.fi => + islossless F.f => islossless D(F, P).distinguish. + + lemma Real_G1 &m: + Pr[GReal(D).main() @ &m: res /\ C.c <= max_size] <= + Pr[G1(DRestr(D)).main() @ &m: res] + + (max_size ^ 2 - max_size)%r * inv 2%r * mu dstate (pred1 witness) + + Pr[G1(DRestr(D)).main() @&m: G1.bcol] + + Pr[G1(DRestr(D)).main() @&m: G1.bext]. + proof. + apply (@RealOrder.ler_trans _ _ _ (Real_Concrete D _ &m))=>//=. + + exact: D_ll. + have : Pr[CF(DRestr(D)).main() @ &m : res] <= + Pr[G1(DRestr(D)).main() @ &m : res] + + Pr[G1(DRestr(D)).main() @ &m : G1.bcol \/ G1.bext]. + + by byequiv (CF_G1 D D_ll)=>//=/#. + have/#:Pr[G1(DRestr(D)).main() @ &m : G1.bcol \/ G1.bext] + <= Pr[G1(DRestr(D)).main() @&m: G1.bcol] + + Pr[G1(DRestr(D)).main() @&m: G1.bext]. + by rewrite Pr[mu_or];smt(Distr.mu_bounded). + qed. + +end section. diff --git a/sha3/proof/smart_counter/SLCommon.ec b/sha3/proof/smart_counter/SLCommon.ec new file mode 100644 index 0000000..fa82181 --- /dev/null +++ b/sha3/proof/smart_counter/SLCommon.ec @@ -0,0 +1,1073 @@ +(** This is a theory for the Squeezeless sponge: where the ideal + functionality is a fixed-output-length random oracle whose output + length is the input block size. We prove its security even when + padding is not prefix-free. **) +require import Core Int Real StdOrder Ring. +require import List FSet SmtMap Common Distr DProd Dexcepted. +require import PROM. + +require (*..*) Indifferentiability. +(*...*) import Capacity IntOrder. + +pragma -oldip. + +(** Really? **) +abbrev ([+]) ['a 'b] (x : 'b) = fun (_ : 'a) => x. + +type state = block * capacity. +op dstate = bdistr `*` cdistr. + +clone include Indifferentiability with + type p <- state, + type f_in <- block list, + type f_out <- block + rename [module] "GReal" as "RealIndif" + [module] "GIdeal" as "IdealIndif". + +(** max number of call to the permutation and its inverse, + including those performed by the construction. *) +op max_size : { int | 0 <= max_size } as max_ge0. + +(** Ideal Functionality **) +clone export Tuple as TupleBl with + type t <- block, + op Support.enum <- Block.blocks + proof Support.enum_spec by exact Block.enum_spec. + +op bl_enum = flatten (mkseq (fun i => wordn i) (max_size + 1)). +op bl_univ = FSet.oflist bl_enum. + +(* -------------------------------------------------------------------------- *) +(* Random oracle from block list to block *) + +clone import FullRO as F with + type in_t <- block list, + type out_t <- block, + op dout _ <- bdistr, + type d_in_t <- unit, + type d_out_t <- bool. +import FullEager. + +module Redo = { + var prefixes : (block list, state) fmap + + proc init() : unit = { + prefixes <- empty.[[] <- (b0,c0)]; + } +}. + +(** We can now define the squeezeless sponge construction **) +module SqueezelessSponge (P:DPRIMITIVE): FUNCTIONALITY = { + proc init () = { + Redo.init(); + } + + proc f(p : block list): block = { + var (sa,sc) <- (b0,c0); + var i : int <- 0; + + while (i < size p) { (* Absorption *) + if (take (i+1) p \in Redo.prefixes) { + (sa,sc) <- oget Redo.prefixes.[take (i+1) p]; + } else { + (sa,sc) <- (sa +^ nth witness p i, sc); + (sa,sc) <@ P.f((sa,sc)); + Redo.prefixes.[take (i+1) p] <- (sa,sc); + } + i <- i + 1; + } + + return sa; (* Squeezing phase (non-iterated) *) + } +}. + +clone export DProd.ProdSampling as Sample2 with + type t1 <- block, + type t2 <- capacity. + +(* -------------------------------------------------------------------------- *) +(** TODO move this **) + +op incl (m m':('a,'b)fmap) = + forall x, m .[x] <> None => m'.[x] = m.[x]. + +(* -------------------------------------------------------------------------- *) +(** usefull type and operators for the proof **) + +type handle = int. + +type hstate = block * handle. + +type ccapacity = capacity * flag. + +type smap = (state , state ) fmap. +type hsmap = (hstate, hstate ) fmap. +type handles = (handle, ccapacity) fmap. + +pred is_pre_permutation (m mi : ('a,'a) fmap) = + (forall x, rng m x => dom mi x) + /\ (forall x, rng mi x => dom m x). + +lemma half_permutation_set (m' mi' : ('a,'a) fmap) x' y': + (forall x, rng m' x => dom mi' x) + => (forall x, rng m'.[x' <- y'] x => dom mi'.[y' <- x'] x). +proof. +move=> h x0; rewrite rngE=> - /= [x]; case: (x = x')=> [<*>|]. ++ by rewrite get_set_sameE=> /= <*>; rewrite domE get_set_sameE. +rewrite get_setE=> -> /= m'x_x0; move: (h x0 _). ++ by rewrite rngE; exists x. +by rewrite mem_set=> ->. +qed. + +lemma pre_permutation_set (m mi : ('a,'a) fmap) x y: + is_pre_permutation m mi => + is_pre_permutation m.[x <- y] mi.[y <- x]. +proof. +move=> [dom_mi dom_m]. +by split; apply/half_permutation_set. +qed. + +(* Functionnal version of the construction using handle *) +op step_hpath (mh:hsmap) (sah:hstate option) (b:block) = + if sah = None + then None + else + let sah = oget sah in + mh.[(sah.`1 +^ b, sah.`2)]. + +op build_hpath (mh:hsmap) (bs:block list) = + foldl (step_hpath mh) (Some (b0,0)) bs. + +inductive build_hpath_spec mh p v h = +| Empty of (p = []) + & (v = b0) + & (h = 0) +| Extend p' b v' h' of (p = rcons p' b) + & (build_hpath mh p' = Some (v',h')) + & (mh.[(v' +^ b,h')] = Some (v,h)). + +lemma build_hpathP mh p v h: + build_hpath mh p = Some (v,h) <=> build_hpath_spec mh p v h. +proof. +elim/last_ind: p v h=> @/build_hpath //= [v h|p b ih v h]. ++ by split=> [!~#] <*>; [exact/Empty|move=> []]; smt(size_rcons size_ge0). +rewrite -{1}cats1 foldl_cat {1}/step_hpath /=. +case: {-1}(foldl _ _ _) (eq_refl (foldl (step_hpath mh) (Some (b0,0)) p))=> //=. ++ apply/implybN; case=> [|p' b0 v' h']. + + smt(size_rcons size_ge0). + move=> ^/rconssI <<- /rconsIs ->>. + by rewrite /build_hpath=> ->. +move=> [v' h']; rewrite -/(build_hpath _ _)=> build. +split. ++ by move=> mh__; apply/(Extend mh (rcons p b) v h p b v' h' _ build mh__). +case=> [| p' b' v'' h'']. ++ smt(size_rcons size_ge0). +move=> ^/rconssI <<- /rconsIs <<-. +by rewrite build /= => [#] <*>. +qed. + +lemma build_hpath_map0 p: + build_hpath empty p = if p = [] then Some (b0,0) else None. +proof. +elim/last_ind: p=> //= p b _. +by rewrite -{1}cats1 /build_hpath foldl_cat {1}/step_hpath /= emptyE /= #smt:(size_rcons size_ge0). +qed. + +(* -------------------------------------------------------------------------- *) +theory Prefix. + +op prefix ['a] (s t : 'a list) = + with s = x :: s', t = y :: t' => if x = y then 1 + prefix s' t' else 0 + with s = _ :: _ , t = [] => 0 + with s = [] , t = _ :: _ => 0 + with s = [] , t = [] => 0. + +lemma prefix0s (s : 'a list): prefix [] s = 0. +proof. by elim: s. qed. + +lemma prefixs0 (s : 'a list): prefix s [] = 0. +proof. by elim: s. qed. + +lemma prefix_eq (l : 'a list) : prefix l l = size l. +proof. elim:l=>//=/#. qed. + + +lemma prefixC (l1 l2 : 'a list) : + prefix l1 l2 = prefix l2 l1. +proof. +move:l1; elim: l2=> //=; first by (move=> l1; elim: l1=> //=). +move=> e2 l2 Hind l1; move: e2 l2 Hind; elim: l1=> //=. +move=> e1 l1 Hind e2 l2 Hind1; rewrite Hind1 /#. +qed. + +lemma prefix_ge0 (l1 l2 : 'a list) : + 0 <= prefix l1 l2. +proof. +move: l2; elim: l1=> //=; first (move=> l2; elim: l2=> //=). +move=> e1 l1 Hind l2; move: e1 l1 Hind; elim: l2=> //=. +move=> e2 l2 Hind2 e1 l1 Hind1 /#. +qed. + +lemma prefix_sizel (l1 l2 : 'a list) : + prefix l1 l2 <= size l1. +proof. +move: l2; elim: l1=> //=; first by (move=> l2; elim: l2=> //=). +move=> e1 l1 Hind l2; move: e1 l1 Hind; elim: l2=> //=; 1:smt(size_ge0). +by move=> e2 l2 Hind2 e1 l1 Hind1; smt(size_ge0). +qed. + +lemma prefix_sizer (l1 l2 : 'a list) : + prefix l1 l2 <= size l2. +proof. +by rewrite prefixC prefix_sizel. +qed. + +lemma prefix_take (l1 l2 : 'a list) : + take (prefix l1 l2) l1 = take (prefix l1 l2) l2. +proof. +move: l2; elim: l1=> //=; first by (move=> l2; elim: l2=> //=). +move=> e1 l1 Hind l2 /=; move: e1 l1 Hind; elim: l2=> //=. +move=> e2 l2 Hind1 e1 l1 Hind2=> //=. +by case: (e1 = e2)=> [-> /#|]. +qed. + +lemma take_take (l : 'a list) (i j : int) : + take i (take j l) = take (min i j) l. +proof. +case: (i <= j)=> Hij. ++ case: (j < size l)=> Hjsize; last smt(take_oversize). + case: (0 <= i)=> Hi0; last smt(take_le0). + apply: (eq_from_nth witness); 1:smt(size_take). + move=> k; rewrite !size_take //= 1:/# Hjsize /=. + have ->: (if i < j then i else j) = i by smt(). + move=> [Hk0 Hki]. + by rewrite !nth_take /#. +case: (0 < j)=> //= Hj0; last smt(take_le0). +rewrite (: min i j = j) 1:minrE 1:/#. +by rewrite take_oversize //= size_take /#. +qed. + +lemma prefix_take_leq (l1 l2 : 'a list) (i : int) : + i <= prefix l1 l2 => take i l1 = take i l2. +proof. +move=> Hi; have ->: i = min i (prefix l1 l2) by smt(minrE). +by rewrite -(take_take l1 i _) -(take_take l2 i _) prefix_take. +qed. + +lemma prefix_nth (l1 l2 : 'a list) : + let i = prefix l1 l2 in + forall j, 0 <= j < i => + nth witness l1 j = nth witness l2 j. +proof. +move=> /=; have Htake:= prefix_take l1 l2. +move=> j [Hj0 Hjp]; rewrite -(nth_take witness (prefix l1 l2)) 1:prefix_ge0 //. +by rewrite -(nth_take witness (prefix l1 l2) l2) 1:prefix_ge0 // Htake. +qed. + +(* TODO: can we define this as a fold on a set instead of on a list? *) +op max_prefix (l1 l2 : 'a list) (ll : 'a list list) = + with ll = "[]" => l2 + with ll = (::) l' ll' => + if prefix l1 l2 < prefix l1 l' then max_prefix l1 l' ll' + else max_prefix l1 l2 ll'. + +op get_max_prefix (l : 'a list) (ll : 'a list list) = + with ll = "[]" => [] + with ll = (::) l' ll' => max_prefix l l' ll'. + +pred prefix_inv (queries : (block list, block) fmap) + (prefixes : (block list, state) fmap) = + (forall (bs : block list), + bs \in queries => oget queries.[bs] = (oget prefixes.[bs]).`1) && + (forall (bs : block list), + bs \in queries => forall i, take i bs \in prefixes) && + (forall (bs : block list), + forall i, take i bs <> [] => + take i bs \in prefixes => + exists l2, (take i bs) ++ l2 \in queries). + +pred all_prefixes (prefixes : (block list, state) fmap) = + forall (bs : block list), bs \in prefixes => forall i, take i bs \in prefixes. + +lemma aux_mem_get_max_prefix (l1 l2 : 'a list) ll : + max_prefix l1 l2 ll = l2 \/ max_prefix l1 l2 ll \in ll. +proof. +move: l1 l2; elim: ll=> //= l3 ll Hind l1 l2. +case: (prefix l1 l2 < prefix l1 l3)=> //= hmax. ++ by have /#:= Hind l1 l3. +by have /#:= Hind l1 l2. +qed. + +lemma mem_get_max_prefix (l : 'a list) ll : + ll <> [] => get_max_prefix l ll \in ll. +proof. +move: l; elim: ll=> //= l2 ll Hind l1. +exact/aux_mem_get_max_prefix. +qed. + +lemma take_get_max_prefix l (prefixes : (block list,state) fmap) : + (exists b, b \in prefixes) => + all_prefixes prefixes => + take (prefix l (get_max_prefix l (elems (fdom prefixes)))) l \in prefixes. +proof. +move=> nil_in_dom all_pref. +rewrite prefix_take all_pref -mem_fdom memE mem_get_max_prefix; smt(memE mem_fdom). +qed. + +lemma take_get_max_prefix2 l (prefixes : (block list,state) fmap) i : + (exists b, b \in prefixes) => + all_prefixes prefixes => + i <= prefix l (get_max_prefix l (elems (fdom prefixes))) => + take i l \in prefixes. +proof. +move=> nil_in_dom all_pref hi. +rewrite (prefix_take_leq _ _ i hi) all_pref -mem_fdom memE mem_get_max_prefix. +smt(memE mem_fdom). +qed. + +lemma prefix_cat (l l1 l2 : 'a list) : + prefix (l ++ l1) (l ++ l2) = size l + prefix l1 l2. +proof. by move: l1 l2; elim: l=> /#. qed. + +lemma prefix_leq_take (l1 l2 : 'a list) i : + 0 <= i <= min (size l1) (size l2) => + take i l1 = take i l2 => + i <= prefix l1 l2. +proof. +move=> [hi0 himax] htake. +rewrite -(cat_take_drop i l1) -(cat_take_drop i l2) htake. +rewrite prefix_cat size_take //=; smt(prefix_ge0). +qed. + +lemma prefix0 (l1 l2 : 'a list) : + prefix l1 l2 = 0 <=> l1 = [] \/ l2 = [] \/ head witness l1 <> head witness l2 . +proof. +move: l2; elim: l1=> //= [[] //=|]. +move=> e1 l1 Hind l2; move: e1 l1 Hind; elim: l2=> //= e2 l2 Hind2 e1 l1 Hind1. +smt(prefix_ge0). +qed. + +lemma head_nth0 (l : 'a list) : head witness l = nth witness l 0. +proof. by elim: l. qed. + +lemma get_prefix (l1 l2 : 'a list) i : + 0 <= i <= min (size l1) (size l2)=> + (drop i l1 = [] \/ drop i l2 = [] \/ + (i < min (size l1) (size l2) /\ + nth witness l1 i <> nth witness l2 i)) => + take i l1 = take i l2 => + i = prefix l1 l2. +proof. +move=>[hi0 hisize] [|[]]. ++ move=> hi. + have:= size_eq0 (drop i l1); rewrite {2}hi /= size_drop // => h. + have hsize: size l1 = i by smt(). + rewrite -hsize take_size. + rewrite -{2}(cat_take_drop (size l1) l2)=> <-. + by rewrite -{2}(cats0 l1) prefix_cat; case: (drop (size l1) l2). ++ move=> hi. + have:= size_eq0 (drop i l2); rewrite {2}hi /= size_drop // => h. + have hsize: size l2 = i by rewrite /#. + rewrite -hsize take_size. + rewrite -{2}(cat_take_drop (size l2) l1)=> ->. + by rewrite -{4}(cats0 l2) prefix_cat; case: (drop (size l2) l1). +move=> [himax hnth] htake. +rewrite -(cat_take_drop i l1) -(cat_take_drop i l2) htake. +rewrite prefix_cat size_take //=. +have [_ ->]:= prefix0 (drop i l1) (drop i l2). ++ case: (i = size l1)=> hi1 //=. + + by rewrite hi1 drop_size //=. + case: (i = size l2)=> hi2 //=. + + by rewrite hi2 drop_size //=. + by rewrite 2!head_nth0 nth_drop //= nth_drop //= hnth. +smt(). +qed. + +lemma get_max_prefix_leq (l1 l2 : 'a list) (ll : 'a list list) : + prefix l1 l2 <= prefix l1 (max_prefix l1 l2 ll). +proof. by move: l1 l2; elim: ll=> /#. qed. + +lemma get_max_prefix_is_max (l1 l2 : 'a list) (ll : 'a list list) : + forall l3, l3 \in ll => prefix l1 l3 <= prefix l1 (max_prefix l1 l2 ll). +proof. +move: l1 l2; elim: ll=> //= l4 ll Hind l1 l2 l3. +by case: (prefix l1 l2 < prefix l1 l4)=> //= h []; smt(get_max_prefix_leq). +qed. + +lemma get_max_prefix_max (l : 'a list) (ll : 'a list list) : + forall l2, l2 \in ll => prefix l l2 <= prefix l (get_max_prefix l ll). +proof. smt(get_max_prefix_is_max get_max_prefix_leq). qed. + +(** TODO: NOT PRETTY! **) +lemma all_take_in (l : block list) i prefixes : + 0 <= i <= size l => + all_prefixes prefixes => + take i l \in prefixes => + i <= prefix l (get_max_prefix l (elems (fdom prefixes))). +proof. +move=>[hi0 hisize] all_prefix take_in_dom. +have ->:i = prefix l (take i l);2:smt(get_max_prefix_max memE mem_fdom). +apply get_prefix. ++ smt(size_take). ++ by right;left;apply size_eq0;rewrite size_drop//size_take//=/#. +smt(take_take). +qed. + +lemma prefix_inv_leq (l : block list) i prefixes queries : + 0 <= i <= size l => + elems (fdom queries) <> [] => + all_prefixes prefixes => + take i l \in prefixes => + prefix_inv queries prefixes => + i <= prefix l (get_max_prefix l (elems (fdom queries))). +proof. +move=>h_i h_nil h_all_prefixes take_in_dom [?[h_prefix_inv h_exist]]. +case(take i l = [])=>//=h_take_neq_nil. ++ smt(prefix_ge0 size_take). +have [l2 h_l2_mem]:=h_exist l i h_take_neq_nil take_in_dom. +rewrite -mem_fdom memE in h_l2_mem. +rewrite(StdOrder.IntOrder.ler_trans _ _ _ _ (get_max_prefix_max _ _ _ h_l2_mem)). +rewrite-{1}(cat_take_drop i l)prefix_cat size_take 1:/#;smt(prefix_ge0). +qed. + + +lemma max_prefix_eq (l : 'a list) (ll : 'a list list) : + max_prefix l l ll = l. +proof. +move:l;elim:ll=>//=l2 ll Hind l1;smt( prefix_eq prefix_sizel). +qed. + +lemma prefix_max_prefix_eq_size (l1 l2 : 'a list) (ll : 'a list list) : + l1 = l2 \/ l1 \in ll => + prefix l1 (max_prefix l1 l2 ll) = size l1. +proof. +move:l1 l2;elim:ll=>//=;1:smt(prefix_eq). +move=>l3 ll Hind l1 l2[->|[->|h1]]. ++ by rewrite prefix_eq max_prefix_eq ltzNge prefix_sizel /= prefix_eq. ++ rewrite prefix_eq max_prefix_eq. + case(prefix l3 l2 < size l3)=>//=h;1:by rewrite prefix_eq. + have h1: prefix l3 l2 = size l3 by smt(prefix_sizel). + have: size l3 <= prefix l3 (max_prefix l3 l2 ll);2:smt(prefix_sizel). + rewrite-h1. + by clear Hind l1 h h1;move:l2 l3;elim:ll=>//=l3 ll Hind l1 l2/#. +by case(prefix l1 l2 < prefix l1 l3)=>//=/#. +qed. + +lemma prefix_get_max_prefix_eq_size (l : 'a list) (ll : 'a list list) : + l \in ll => + prefix l (get_max_prefix l ll) = size l. +proof. +move:l;elim:ll=>//;smt(prefix_max_prefix_eq_size). +qed. + +lemma get_max_prefix_exists (l : 'a list) (ll : 'a list list) : + ll <> [] => + exists l2, take (prefix l (get_max_prefix l ll)) l ++ l2 \in ll. +proof. +move:l;elim:ll=>//=l2 ll Hind l1;clear Hind;move:l1 l2;elim:ll=>//=. ++ smt(cat_take_drop prefix_take). +move=>l3 ll Hind l1 l2. +case( prefix l1 l2 < prefix l1 l3 )=>//=h/#. +qed. + +lemma prefix_geq (l1 l2 : 'a list) : + prefix l1 l2 = prefix (take (prefix l1 l2) l1) (take (prefix l1 l2) l2). +proof. +move:l2;elim:l1=>//=[[] //=|] e1 l1 Hind l2;elim:l2=>//=e2 l2 Hind2. +case(e1=e2)=>//=h12. +have ->/=:! 1 + prefix l1 l2 <= 0 by smt(prefix_ge0). +rewrite h12/=/#. +qed. + +lemma prefix_take_prefix (l1 l2 : 'a list) : + prefix (take (prefix l1 l2) l1) l2 = prefix l1 l2. +proof. +move:l2;elim:l1=>//=e1 l1 Hind l2;elim:l2=>//=e2 l2 Hind2. +case(e1=e2)=>//=h12. +have ->/=:! 1 + prefix l1 l2 <= 0 by smt(prefix_ge0). +rewrite h12/=/#. +qed. + +lemma prefix_leq_prefix_cat (l1 l2 l3 : 'a list) : + prefix l1 l2 <= prefix (l1 ++ l3) l2. +proof. +move:l2 l3;elim l1=>//= [[]|]; 1,2:smt(take_le0 prefix_ge0). +move=>e1 l1 hind1 l2;elim:l2=>//=e2 l2 hind2 l3/#. +qed. + +lemma prefix_take_leq_prefix (l1 l2 : 'a list) i : + prefix (take i l1) l2 <= prefix l1 l2. +proof. +rewrite-{2}(cat_take_drop i l1). +move:(take i l1)(drop i l1);clear i l1=>l1 l3. +exact prefix_leq_prefix_cat. +qed. + +lemma prefix_take_geq_prefix (l1 l2 : 'a list) i : + prefix l1 l2 <= i => + prefix l1 l2 = prefix (take i l1) l2. +proof. +move=>hi. +have: prefix (take i l1) l2 <= prefix l1 l2. ++ rewrite-{2}(cat_take_drop i l1) prefix_leq_prefix_cat. +have /#: prefix l1 l2 <= prefix (take i l1) l2. +rewrite -prefix_take_prefix. +rewrite -(cat_take_drop (prefix l1 l2) (take i l1))take_take minrE hi //=. +by rewrite prefix_leq_prefix_cat. +qed. + +lemma get_max_prefix_take (l : 'a list) (ll : 'a list list) i : + prefix l (get_max_prefix l ll) <= i => + get_max_prefix l ll = get_max_prefix (take i l) ll. +proof. +move:l;elim:ll=>//=l2 ll Hind l1;clear Hind;move:l1 l2;elim:ll=>//=l3 ll Hind l1 l2. +case( prefix l1 l2 < prefix l1 l3 )=>//=h hi. ++ rewrite -prefix_take_geq_prefix//=;1:smt(get_max_prefix_leq). + rewrite -prefix_take_geq_prefix//=;1:smt(get_max_prefix_leq). + rewrite h/=/#. +rewrite -prefix_take_geq_prefix//=;1:smt(get_max_prefix_leq). +rewrite -prefix_take_geq_prefix//=;1:smt(get_max_prefix_leq). +rewrite h/=/#. +qed. + + +lemma drop_prefix_neq (l1 l2 : 'a list) : + drop (prefix l1 l2) l1 = [] \/ drop (prefix l1 l2) l1 <> drop (prefix l1 l2) l2. +proof. +move: l2; elim: l1=> //= e1 l1 hind1; elim=> //= e2 l2 //= hind2 //=. +smt(prefix_ge0). +qed. + +lemma prefix_prefix_prefix (l1 l2 l3 : 'a list) (ll : 'a list list) : + prefix l1 l2 <= prefix l1 l3 => + prefix l1 (max_prefix l1 l2 ll) <= prefix l1 (max_prefix l1 l3 ll). +proof. +move:l1 l2 l3;elim:ll=>//=l4 ll hind l1 l2 l3 h123/#. +qed. + +lemma prefix_lt_size (l : 'a list) (ll : 'a list list) : + prefix l (get_max_prefix l ll) < size l => + forall i, prefix l (get_max_prefix l ll) < i => + ! take i l \in ll. +proof. +move:l;elim:ll=>//=l2 ll Hind l1;clear Hind;move:l1 l2;elim:ll=>//=. ++ progress. + rewrite-(cat_take_drop (prefix l1 l2) (take i l1)) + -{3}(cat_take_drop (prefix l1 l2) l2)take_take/min H0/=. + rewrite prefix_take. + have: drop (prefix l1 l2) (take i l1) <> drop (prefix l1 l2) l2;2:smt(catsI). + rewrite (prefix_take_geq_prefix l1 l2 i) 1:/#. + have:= drop_prefix_neq (take i l1) l2. + have /#: drop (prefix (take i l1) l2) (take i l1) <> []. + have: 0 < size (drop (prefix (take i l1) l2) (take i l1));2:smt(size_eq0). + rewrite size_drop 1:prefix_ge0 size_take;1:smt(prefix_ge0). + by rewrite-prefix_take_geq_prefix /#. + +move=>l3 ll hind l1 l2. +case(prefix l1 l2 < prefix l1 l3)=>//=h;progress. ++ rewrite!negb_or/=. + have:= hind l1 l3 H i H0;rewrite negb_or=>[][->->]/=. + have:= hind l1 l2 _ i _;smt(prefix_prefix_prefix). +smt(prefix_prefix_prefix). +qed. + +lemma asfadst queries prefixes (bs : block list) : + prefix_inv queries prefixes => + elems (fdom queries ) <> [] => + all_prefixes prefixes => + (forall j, 0 <= j <= size bs => take j bs \in prefixes) => + take (prefix bs (get_max_prefix bs (elems (fdom queries))) + 1) bs = bs. +proof. +progress. +have h:=prefix_inv_leq bs (size bs) prefixes queries _ _ _ _ _;rewrite//=. ++ rewrite H2//=;exact size_ge0. +have ->/=: prefix bs (get_max_prefix bs (elems (fdom queries))) = size bs by smt(prefix_sizel). +rewrite take_oversize/#. +qed. + + +lemma prefix_exchange_prefix_inv (ll1 ll2 : 'a list list) (l : 'a list) : + (forall l2, l2 \in ll1 => l2 \in ll2) => + (forall (l2 : 'a list), l2 \in ll1 => forall i, take i l2 \in ll2) => + (forall l2, l2 \in ll2 => exists l3, l2 ++ l3 \in ll1) => + prefix l (get_max_prefix l ll1) = prefix l (get_max_prefix l ll2). +proof. +case(ll1 = [])=>//=[-> _ _|]. ++ by case: (ll2 = [])=> [->> //=|] //= + /mem_eq0. +move=> ll1_nil incl all_prefix incl2; have ll2_nil: ll2 <> [] by smt(mem_eq0). +have:= get_max_prefix_max l ll2 (get_max_prefix l ll1) _. ++ by rewrite incl mem_get_max_prefix ll1_nil. +have mem_ll2:=mem_get_max_prefix l ll2 ll2_nil. +have[]l3 mem_ll1:=incl2 _ mem_ll2. +have:=get_max_prefix_max l ll1 _ mem_ll1. +smt(prefixC prefix_leq_prefix_cat). +qed. + +lemma prefix_inv_nil queries prefixes : + prefix_inv queries prefixes => + elems (fdom queries) = [] => fdom prefixes \subset fset1 []. +proof. +move=>[h1 [h2 h3]] h4 x h5;rewrite in_fset1. +have:=h3 x (size x). +rewrite take_size -mem_fdom h5/=;apply absurd=>//=h6. +rewrite h6/=negb_exists/=;smt(memE mem_fdom). +qed. + +lemma aux_prefix_exchange queries prefixes (l : block list) : + prefix_inv queries prefixes => all_prefixes prefixes => + elems (fdom queries) <> [] => + prefix l (get_max_prefix l (elems (fdom queries))) = + prefix l (get_max_prefix l (elems (fdom prefixes))). +proof. +move=>[h1[h2 h3]] h5 h4;apply prefix_exchange_prefix_inv. ++ move=> l2; rewrite -memE mem_fdom=> /h2 /(_ (size l2)). + by rewrite take_size -mem_fdom memE. ++ move=> l2; rewrite -memE mem_fdom=> /h2 + i - /(_ i). + by rewrite -mem_fdom memE. +move=>l2; rewrite -memE=> mem_l2. +case(l2=[])=>//=hl2;1:rewrite hl2/=. ++ move:h4;apply absurd=>//=;rewrite negb_exists/= => /mem_eq0 //=. +have:= h3 l2 (size l2); rewrite take_size hl2 -mem_fdom mem_l2. +by move=> /= [] l3 hl3; exists l3; rewrite -memE mem_fdom. +qed. + +lemma prefix_exchange queries prefixes (l : block list) : + prefix_inv queries prefixes => all_prefixes prefixes => + prefix l (get_max_prefix l (elems (fdom queries))) = + prefix l (get_max_prefix l (elems (fdom prefixes))). +proof. +move=> [h1[h2 h3]] h5. +case: (elems (fdom queries) = [])=> h4. ++ have h6:=prefix_inv_nil queries prefixes _ h4;1:rewrite/#. + rewrite h4/=. + have fdom_prefixP: fdom prefixes = fset0 \/ fdom prefixes = fset1 []. + + by move: h6; rewrite !fsetP /(\subset); smt(in_fset0 in_fset1). + case(elems (fdom prefixes) = [])=>//=[->//=|]h7. + have h8:elems (fdom prefixes) = [[]]. + + have []:= fdom_prefixP. + + by move=> h8; move: h7; rewrite h8 elems_fset0. + by move=> ->; rewrite elems_fset1. + by rewrite h8=>//=. +by apply/(aux_prefix_exchange _ _ _ _ h5 h4). +qed. + +pred all_prefixes_fset (prefixes : block list fset) = + forall bs, bs \in prefixes => forall i, take i bs \in prefixes. + +pred inv_prefix_block (queries : (block list, block) fmap) + (prefixes : (block list, block) fmap) = + (forall (bs : block list), + bs \in queries => queries.[bs] = prefixes.[bs]) && + (forall (bs : block list), + bs \in queries => forall i, 0 < i <= size bs => take i bs \in prefixes). + +lemma prefix_gt0_mem l (ll : 'a list list) : + 0 < prefix l (get_max_prefix l ll) => + get_max_prefix l ll \in ll. +proof. +move:l;elim:ll=>//=;first by move=>l;elim:l. +move=>l2 ll hind l1;clear hind;move:l1 l2;elim:ll=>//=l3 ll hind l1 l2. +by case(prefix l1 l2 < prefix l1 l3)=>//=/#. +qed. + +lemma inv_prefix_block_mem_take queries prefixes l i : + inv_prefix_block queries prefixes => + 0 < i < prefix l (get_max_prefix l (elems (fdom queries))) => + take i l \in prefixes. +proof. +move=>[]H_incl H_all_prefixes Hi. +rewrite (prefix_take_leq _ (get_max_prefix l (elems (fdom queries))))1:/#. +rewrite H_all_prefixes. +have:get_max_prefix l (elems (fdom queries)) \in queries;2:smt(domE). +by rewrite -mem_fdom memE;apply prefix_gt0_mem=>/#. +smt(prefix_sizer). +qed. + +lemma prefix_cat_leq_prefix_size (l1 l2 l3 : 'a list): + prefix (l1 ++ l2) l3 <= prefix l1 l3 + size l2. +proof. +move:l2 l3;elim:l1=>//=. ++ by move=> l2 []; smt(prefix_sizel). +by move=>e1 l1 hind1 l2 l3;move:e1 l1 l2 hind1;elim:l3=>//=;1:smt(size_ge0). +qed. + +lemma prefix_cat1 (l1 l2 l3 : 'a list) : + prefix (l1 ++ l2) l3 = prefix l1 l3 + + if prefix l1 l3 = size l1 + then prefix l2 (drop (size l1) l3) + else 0. +proof. +move:l2 l3;elim:l1=>//=. ++ by move=> l2 []; smt(prefix_sizel). +move=>e1 l1 hind1 l2 l3;move:e1 l1 l2 hind1;elim:l3=>//=;1:smt(size_ge0). +by move=>e3 l3 hind3 e1 l1 l2 hind1;case(e1=e3)=>//=[->>|h];smt(size_ge0). +qed. + + +lemma prefix_leq_prefix_cat_size (l1 l2 : 'a list) (ll : 'a list list) : + prefix (l1++l2) (get_max_prefix (l1++l2) ll) <= + prefix l1 (get_max_prefix l1 ll) + + if (prefix l1 (get_max_prefix l1 ll) = size l1) + then prefix l2 (get_max_prefix l2 (map (drop (size l1)) ll)) + else 0. +proof. +move:l1 l2;elim:ll=>//=. ++ smt(prefixs0). +move=>l3 ll hind{hind};move:l3;elim:ll=>//=;1:smt(prefix_cat1). +move=>l4 ll hind l3 l1 l2. +case(prefix (l1 ++ l2) l3 < prefix (l1 ++ l2) l4)=>//=. ++ rewrite 2!prefix_cat1. + case(prefix l1 l3 = size l1)=>//=H_l1l3;case(prefix l1 l4 = size l1)=>//=H_l1l4. + - rewrite H_l1l4 H_l1l3/=ltz_add2l=>h;rewrite h/=. + rewrite(StdOrder.IntOrder.ler_trans _ _ _ (hind _ _ _)). + have->/=:prefix l1 (max_prefix l1 l4 ll) = size l1 + by move:{hind};elim:ll=>//=;smt(prefix_sizel). + by have->/=:prefix l1 (max_prefix l1 l3 ll) = size l1 + by move:{hind};elim:ll=>//=;smt(prefix_sizel). + - smt(prefix_sizel prefix_ge0). + - have->/=h:prefix l1 l3 < prefix l1 l4 by smt(prefix_sizel). + rewrite(StdOrder.IntOrder.ler_trans _ _ _ (hind _ _ _)). + have->/=:prefix l1 (max_prefix l1 l4 ll) = size l1 + by move:{hind};elim:ll=>//=;smt(prefix_sizel). + smt(prefix_prefix_prefix). + move=>H_l3l4;rewrite H_l3l4/=. + rewrite(StdOrder.IntOrder.ler_trans _ _ _ (hind _ _ _)). + by case(prefix l1 (max_prefix l1 l4 ll) = size l1)=>//=->; + smt(prefix_prefix_prefix). +rewrite 2!prefix_cat1. +case(prefix l1 l3 = size l1)=>//=H_l1l3;case(prefix l1 l4 = size l1)=>//=H_l1l4. ++ by rewrite H_l1l4 H_l1l3/=ltz_add2l=>h;rewrite h/=hind. ++ rewrite H_l1l3. + have->/=:!size l1 < prefix l1 l4 by smt(prefix_sizel). + rewrite(StdOrder.IntOrder.ler_trans _ _ _ (hind _ _ _))//=. + have->//=:prefix l1 (max_prefix l1 l3 ll) = size l1 + by move:{hind};elim:ll=>//=;smt(prefix_sizel). + smt(prefix_prefix_prefix). ++ smt(prefix_sizel prefix_ge0). +move=>H_l3l4;rewrite H_l3l4/=. +rewrite(StdOrder.IntOrder.ler_trans _ _ _ (hind _ _ _))//=. +smt(prefix_prefix_prefix). +qed. + + +lemma diff_size_prefix_leq_cat (l1 l2 : 'a list) (ll : 'a list list) : + size l1 - prefix l1 (get_max_prefix l1 ll) <= + size (l1++l2) - prefix (l1++l2) (get_max_prefix (l1++l2) ll). +proof. +smt(prefix_leq_prefix_cat_size prefix_sizel prefix_ge0 size_ge0 prefix_sizer size_cat). +qed. + + + +(* lemma prefix_inv_prefix queries prefixes l : *) +(* prefix_inv queries prefixes => *) +(* all_prefixes prefixes => *) +(* (elems (fdom queries) = [] => elems (fdom prefixes) = [[]]) => *) +(* prefix l (get_max_prefix l (elems (fdom queries))) = *) +(* prefix l (get_max_prefix l (elems (fdom prefixes))). *) +(* proof. *) +(* move=>[? h_prefix_inv] h_all_prefixes. *) +(* case(elems (fdom queries) = [])=>//=h_nil. *) +(* + by rewrite h_nil//==>->/=. *) +(* cut h_mem_queries:=mem_get_max_prefix l (elems (fdom queries)) h_nil. *) +(* cut h_leq :=all_take_in l (prefix l (get_max_prefix l (elems (fdom queries)))) _ _ h_all_prefixes _. *) +(* + smt(prefix_ge0 prefix_sizel). *) +(* + by rewrite prefix_take h_prefix_inv memE h_mem_queries. *) +(* cut:=all_take_in l (prefix l (get_max_prefix l (elems (fdom prefixes)))) _ _ h_all_prefixes _. *) +(* + smt(prefix_ge0 prefix_sizel). *) +(* + *) +(* rewrite prefix_take. *) + +(* rewrite -take_size. *) + +(* print mem_get_max_prefix. *) + +(* qed. *) + +pred invm (m mi : ('a * 'b, 'a * 'b) fmap) = + forall x y, m.[x] = Some y <=> mi.[y] = Some x. + +lemma invm_set (m mi : ('a * 'b, 'a * 'b) fmap) x y : + ! x \in m => ! rng m y => invm m mi => invm m.[x <- y] mi.[y <- x]. +proof. +move=>Hxdom Hyrng Hinv a b; rewrite !get_setE; split. ++ case(a=x)=>//=hax hab;have->/#:b<>y. + by have/#: rng m b;rewrite rngE /#. +case(a=x)=>//=hax. ++ case(b=y)=>//=hby. + by rewrite (eq_sym y b)hby/=-Hinv hax;rewrite domE /=/# in Hxdom. +by rewrite Hinv/#. +qed. + +end Prefix. +export Prefix. + +(* -------------------------------------------------------------------------- *) + +module C = { + var c : int + var queries : (block list, block) fmap + proc init () = { + c <- 0; + queries <- empty.[[] <- b0]; + } +}. + +module PC (P:PRIMITIVE) = { + + proc init () = { + C.init(); + P.init(); + } + + proc f (x:state) = { + var y <- (b0,c0); + y <@ P.f(x); + C.c <- C.c + 1; + return y; + } + + proc fi(x:state) = { + var y <- (b0,c0); + y <@ P.fi(x); + C.c <- C.c + 1; + return y; + } + +}. + +module DPRestr (P:DPRIMITIVE) = { + + proc f (x:state) = { + var y <- (b0,c0); + if (C.c + 1 <= max_size) { + y <@ P.f(x); + C.c <- C.c + 1; + } + return y; + } + + proc fi(x:state) = { + var y <- (b0,c0); + if (C.c + 1 <= max_size) { + y <@ P.fi(x); + C.c <- C.c + 1; + } + return y; + } + +}. + +module PRestr (P:PRIMITIVE) = { + + proc init () = { + C.init(); + P.init(); + } + + proc f = DPRestr(P).f + + proc fi = DPRestr(P).fi + +}. + +module FC(F:FUNCTIONALITY) = { + + proc init() = { + F.init(); + } + + proc f (bs:block list) = { + var b <- b0; + if (bs \notin C.queries) { + C.c <- C.c + size bs - prefix bs (get_max_prefix bs (elems (fdom C.queries))); + b <@ F.f(bs); + C.queries.[bs] <- b; + } else { + b <- oget C.queries.[bs]; + } + return b; + } +}. + +module DFRestr(F:DFUNCTIONALITY) = { + + proc f (bs:block list) = { + var b <- b0; + if (bs \notin C.queries) { + if (C.c + size bs - prefix bs (get_max_prefix bs (elems (fdom C.queries))) <= max_size) { + C.c <- C.c + size bs - prefix bs (get_max_prefix bs (elems (fdom C.queries))); + b <@ F.f(bs); + C.queries.[bs] <- b; + } + } else { + b <- oget C.queries.[bs]; + } + return b; + } +}. + +module FRestr(F:FUNCTIONALITY) = { + + proc init() = { + Redo.init(); + F.init(); + } + + proc f = DFRestr(F).f + +}. + +(* -------------------------------------------------------------------------- *) +(* This allow swap the counting from oracle to adversary *) +module DRestr(D:DISTINGUISHER, F:DFUNCTIONALITY, P:DPRIMITIVE) = { + proc distinguish() = { + var b; + C.init(); + b <@ D(DFRestr(F), DPRestr(P)).distinguish(); + return b; + } +}. + +lemma rp_ll (P<:DPRIMITIVE{-C}): islossless P.f => islossless DPRestr(P).f. +proof. move=>Hll;proc;sp;if;auto;call Hll;auto. qed. + +lemma rpi_ll (P<:DPRIMITIVE{-C}): islossless P.fi => islossless DPRestr(P).fi. +proof. move=>Hll;proc;sp;if;auto;call Hll;auto. qed. + +lemma rf_ll (F<:DFUNCTIONALITY{-C}): islossless F.f => islossless DFRestr(F).f. +proof. move=>Hll;proc;sp;if;auto;if=>//;auto;call Hll;auto. qed. + +lemma DRestr_ll (D<:DISTINGUISHER{-C}): + (forall (F<:DFUNCTIONALITY{-D})(P<:DPRIMITIVE{-D}), + islossless P.f => islossless P.fi => islossless F.f => + islossless D(F,P).distinguish) => + forall (F <: DFUNCTIONALITY{-DRestr(D)}) (P <: DPRIMITIVE{-DRestr(D)}), + islossless P.f => + islossless P.fi => islossless F.f => islossless DRestr(D, F, P).distinguish. +proof. + move=> D_ll F P p_ll pi_ll f_ll;proc. + call (D_ll (DFRestr(F)) (DPRestr(P)) _ _ _). + + by apply (rp_ll P). + by apply (rpi_ll P). + by apply (rf_ll F). + by inline *;auto. +qed. + +section RESTR. + + declare module F <: FUNCTIONALITY{-C}. + declare module P <: PRIMITIVE{-C, -F}. + declare module D <: DISTINGUISHER{-F, -P, -C}. + + lemma swap_restr &m: + Pr[Indif(FRestr(F), PRestr(P), D).main()@ &m: res] = + Pr[Indif(F,P,DRestr(D)).main()@ &m: res]. + proof. + byequiv=>//;auto. + proc;inline *;wp. + swap{1}[1..2] 3;sim;auto;call(:true);auto. + qed. + +end section RESTR. + +section COUNT. + + declare module P <: PRIMITIVE{-C}. + declare module CO <: CONSTRUCTION{-C, -P}. + declare module D <: DISTINGUISHER{-C, -P, -CO}. + + declare axiom f_ll : islossless P.f. + declare axiom fi_ll : islossless P.fi. + + declare axiom CO_ll : islossless CO(P).f. + + declare axiom D_ll (F <: DFUNCTIONALITY{-D}) (P <: DPRIMITIVE{-D}): + islossless P.f => islossless P.fi => islossless F.f => + islossless D(F, P).distinguish. + + lemma Pr_restr &m : + Pr[Indif(FC(CO(P)), PC(P), D).main()@ &m:res /\ C.c <= max_size] <= + Pr[Indif(CO(P), P, DRestr(D)).main()@ &m:res]. + proof. + byequiv (_: ={glob D, glob P, glob CO} ==> C.c{1} <= max_size => ={res})=>//; + 2:by move=> ??H[]?/H<-. + symmetry;proc;inline *;wp. + call (_: max_size < C.c, ={glob P, glob CO, glob C}). + + apply D_ll. + + proc; sp;if{1};1:by auto;call(_:true);auto. + by auto;call{2} f_ll;auto=>/#. + + by move=> ?_;proc;sp;auto;if;auto;call f_ll;auto. + + by move=> _;proc;sp;auto;call f_ll;auto=>/#. + + proc;sp;auto;if{1};1:by auto;call(_:true);auto. + by call{2} fi_ll;auto=>/#. + + by move=> ?_;proc;sp;auto;if;auto;call fi_ll;auto. + + by move=> _;proc;sp;auto;call fi_ll;auto=>/#. + + proc;inline*;sp 1 1;if;auto;if{1};auto;1:by call(_: ={glob P});auto;sim. + by call{2} CO_ll;auto=>/#. + + by move=> ?_;proc;sp;if;auto;if;auto;call CO_ll;auto. + + by move=> _;proc;sp;if;auto;call CO_ll;auto;smt(prefix_sizel). + auto;call (_:true);auto;call(:true);auto=>/#. + qed. + +end section COUNT. + +(** Operators and properties of handles *) +op hinv (handles:handles) (c:capacity) = + find (fun _ => pred1 c \o fst) handles. + +op hinvK (handles:handles) (c:capacity) = + find (fun _ => pred1 c) (restr Known handles). + +op huniq (handles:handles) = + forall h1 h2 cf1 cf2, + handles.[h1] = Some cf1 => + handles.[h2] = Some cf2 => + cf1.`1 = cf2.`1 => h1 = h2. + +lemma hinvP handles c: + if hinv handles c = None then forall h f, handles.[h] <> Some(c,f) + else exists f, handles.[oget (hinv handles c)] = Some(c,f). +proof. +move=> @/hinv. +have @/pred1 @/(\o) /> [-> /= + h f|h [] /> f -> //= Hmem] := findP (fun _=> pred1 c \o fst) handles. ++ by move=> /(_ h); rewrite domE; case: (handles.[h])=> /#. +by exists f. +qed. + +lemma huniq_hinv (handles:handles) (h:handle): + huniq handles => dom handles h => hinv handles (oget handles.[h]).`1 = Some h. +proof. +move=> Huniq;pose c := (oget handles.[h]).`1. +have:=Huniq h;have:=hinvP handles c. +case (hinv _ _)=> /=[Hdiff _| h' +/(_ h')]. ++ rewrite domE /=; move: (Hdiff h (oget handles.[h]).`2). + by rewrite /c; case: handles.[h]=> //= - []. +move=> [f ->] /(_ (oget handles.[h]) (c,f)) H1 H2;rewrite H1 //. +by move: H2; rewrite domE; case: (handles.[h]). +qed. + +lemma hinvKP handles c: + if hinvK handles c = None then forall h, handles.[h] <> Some(c,Known) + else handles.[oget (hinvK handles c)] = Some(c,Known). +proof. +rewrite /hinvK. +have @/pred1 /= [-> /= + h|h /> -> /=]:= findP (+ pred1 c) (restr Known handles). ++ by move=> /(_ h); rewrite domE restrP=> /#. +by rewrite restrP; case: (handles.[h])=> //= - [] /#. +qed. + +lemma huniq_hinvK (handles:handles) c: + huniq handles => rng handles (c,Known) => handles.[oget (hinvK handles c)] = Some(c,Known). +proof. +move=> Huniq;rewrite rngE=> -[h]H;case: (hinvK _ _) (Huniq h) (hinvKP handles c)=>//=. +by move=>_/(_ h);rewrite H. +qed. + +lemma huniq_hinvK_h h (handles:handles) c: + huniq handles => handles.[h] = Some (c,Known) => hinvK handles c = Some h. +proof. +by move=> Huniq;case: (hinvK _ _) (hinvKP handles c)=>/= [ H | h' /Huniq H/H //]; apply H. +qed. + +(* -------------------------------------------------------------------------- *) +(** The initial Game *) +module GReal(D:DISTINGUISHER) = RealIndif(SqueezelessSponge, PC(Perm), D).