diff --git a/src/.gitignore b/.gitignore similarity index 62% rename from src/.gitignore rename to .gitignore index de6e398..df1285d 100644 --- a/src/.gitignore +++ b/.gitignore @@ -1,3 +1,5 @@ +.nia.cache +Makefile.coq *.vo *.v.d *.glob diff --git a/Makefile b/Makefile new file mode 100644 index 0000000..de5cce5 --- /dev/null +++ b/Makefile @@ -0,0 +1,16 @@ +.PHONY: coq +coq: Makefile.coq + make -f Makefile.coq + +Makefile.coq: Makefile _CoqProject + coq_makefile -f _CoqProject > $@ + +SORT_COQPROJECT = sed 's,[^/]*/,~&,g' | env LC_COLLATE=C sort | sed 's,~,,g' +.PHONY: update-_CoqProject +update-_CoqProject:: + (echo '-R src ""'; (git ls-files 'src/*.v' | $(SORT_COQPROJECT))) > _CoqProject + +.PHONY: clean +clean: + make -f Makefile.coq clean || true + rm -f Makefile.coq || true diff --git a/README.md b/README.md index db5b323..4095121 100644 --- a/README.md +++ b/README.md @@ -1,12 +1,37 @@ - Foundational Cryptography Framework for machine-checked proofs of cryptography in the computational model. -Getting Started: -1) cd src/FCF +# Dependencies + +Coq 8.6 or 8.5 + +# Building -2) make +`make` + +# Exploring Then open a simple example from the “src/examples” directory in Proof General. A good place to start is “ElGamal.v.” Interactively step through this proof to learn how to develop a simple proof of non-adaptive security in the concrete setting. PRF_Encryption_IND_CPA.v contains a more complex proof of adaptive security along with a proof in the asymptotic setting. + The publications describing FCF are available at adam.petcher.net. +# Importing + + env COQPATH=/path/to/fcf/src proofgeneral test.v + Require Import FCF.FCF. + +# History + +This repository used to contain a proof of security of ESPADA SSE Scheme under +`src/ESPADA`. This proof is no longer maintained; it is preserved in git +history. + +Some files that no longer fully build are preserved under `src/FCF/Broken`. +There be dragons, though -- not all of them were ever finished, and some contain +significant `admit`s. + +# Ackowledgements + +This work is sponsored by the Department of the Air Force under Air Force Contract FA8721-05-C-0002. Opinions, interpretations, conclusions, and recommendations are those of the author and are not necessarily endorsed by the United States Government. + +This work is sponsored by the Intelligence Advanced Research Projects Activity under Air Force Contract FA8721-05-C-0002. Opinions, interpretations, conclusions, and recommendations are those of the author and are not necessarily endorsed by the United States Government. diff --git a/_CoqProject b/_CoqProject new file mode 100644 index 0000000..29d0329 --- /dev/null +++ b/_CoqProject @@ -0,0 +1,74 @@ +-R src "" +src/FCF/Admissibility.v +src/FCF/Array.v +src/FCF/Asymptotic.v +src/FCF/Bernoulli.v +src/FCF/Blist.v +src/FCF/Class.v +src/FCF/Comp.v +src/FCF/CompFold.v +src/FCF/ConstructedFunc.v +src/FCF/Crypto.v +src/FCF/DetSem.v +src/FCF/DiffieHellman.v +src/FCF/DistRules.v +src/FCF/DistSem.v +src/FCF/DistTacs.v +src/FCF/Encryption.v +src/FCF/Encryption_PK.v +src/FCF/EqDec.v +src/FCF/ExpectedPolyTime.v +src/FCF/FCF.v +src/FCF/Fold.v +src/FCF/GenTacs.v +src/FCF/GroupTheory.v +src/FCF/HasDups.v +src/FCF/Hybrid.v +src/FCF/Limit.v +src/FCF/Lognat.v +src/FCF/NoDup_gen.v +src/FCF/NotationV1.v +src/FCF/NotationV2.v +src/FCF/OTP.v +src/FCF/OracleCompFold.v +src/FCF/OracleHybrid.v +src/FCF/PRF.v +src/FCF/PRF_Convert.v +src/FCF/PRG.v +src/FCF/PRP_PRF.v +src/FCF/ProgTacs.v +src/FCF/ProgramLogic.v +src/FCF/Rat.v +src/FCF/RepeatCore.v +src/FCF/RndDup.v +src/FCF/RndGrpElem.v +src/FCF/RndInList.v +src/FCF/RndListElem.v +src/FCF/RndNat.v +src/FCF/RndPerm.v +src/FCF/SemEquiv.v +src/FCF/SplitVector.v +src/FCF/StdNat.v +src/FCF/Tactics.v +src/FCF/TwoWorldsEquiv.v +src/FCF/WC_PolyTime.v +src/FCF/WC_PolyTime_old.v +src/FCF/Broken/Encryption_2W.v +src/FCF/Broken/ListHybrid.v +src/FCF/Broken/Procedure.v +src/FCF/Broken/RandPermSwitching.v +src/FCF/Broken/Sigma.v +src/FCF/Broken/State.v +src/FCF/examples/Commit.v +src/FCF/examples/EC_DRBG.v +src/FCF/examples/ElGamal.v +src/FCF/examples/PRF_DRBG.v +src/FCF/examples/PRF_Encryption_IND_CPA.v +src/HMAC/GHMAC_PRF.v +src/HMAC/GNMAC_PRF.v +src/HMAC/HMAC_PRF.v +src/HMAC/HMAC_spec.v +src/HMAC/Message.v +src/HMAC/NMAC_to_HMAC.v +src/HMAC/cAU.v +src/HMAC/hF.v diff --git a/src/ESPADA/.dir-locals.el b/src/ESPADA/.dir-locals.el deleted file mode 100644 index 58c6bff..0000000 --- a/src/ESPADA/.dir-locals.el +++ /dev/null @@ -1 +0,0 @@ -((coq-mode . ((coq-load-path . ("../FCF"))))) diff --git a/src/ESPADA/ESPADA_SSE_OXT.v b/src/ESPADA/ESPADA_SSE_OXT.v deleted file mode 100644 index 3caaf54..0000000 --- a/src/ESPADA/ESPADA_SSE_OXT.v +++ /dev/null @@ -1,12243 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import Coq.ZArith.Zdigits. - -Require Import FCF. -Require Import CompFold. -Require Import GroupTheory. -Require Import RndPerm. -Require Import PRF. - -Require Import SSE. -Require Import TSet. - -Local Open Scope list_scope. -Local Open Scope type_scope. - -Section ESPADA_SSE_OXT. - - Variable lambda : posnat. - Definition DB := DB lambda. - - Variable p : posnat. - Definition multModP (a b : nat) := - modNat (a * b) p. - - Context `{FCG : FiniteCyclicGroup nat}. - - Hypothesis op_multModP : - groupOp = multModP. - - (* PRFs *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - Variable F_P : Bvector lambda -> Blist -> nat. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - Hypothesis Enc_correct : - forall k p c, - In c (getSupport (Enc k p)) -> - Dec k c = p. - - - Instance nz_posnat_plus : - forall (x y : posnat), - nz (x + y). - - intuition. - unfold posnatToNat, natToPosnat. - destruct x. - destruct y. - econstructor. - omega. - - Qed. - - (* TSet *) - Variable TSet : Set. - Hypothesis TSet_EqDec : EqDec TSet. - Variable TSetKey : Set. - Hypothesis TSetKey_EqDec : EqDec TSetKey. - Variable Tag : Set. - Hypothesis Tag_EqDec : EqDec Tag. - Variable TSetSetup : T (pos (lambda + lambda)) -> Comp (TSet * TSetKey). - Variable TSetGetTag : TSetKey -> Keyword -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector (lambda + lambda))). - Variable Leakage_T : Set. - Hypothesis Leakage_T_EqDec : EqDec Leakage_T. - Variable L_T : T (pos (lambda + lambda)) -> list Keyword -> Leakage_T. - - Definition XSet := list nat. - Definition EDB := TSet * XSet. - Definition Key := Bvector lambda * Bvector lambda * Bvector lambda * Bvector lambda * TSetKey. - Definition Query := Keyword * Keyword. - - Definition natToBvector(n : nat) : Bvector lambda := - Z_to_binary lambda (Z.of_nat n). - - Definition natToBlist (n : nat) := - Vector.to_list (natToBvector n). - - Definition bvectorToNat(v : Bvector lambda) : nat := - Z.to_nat (binary_value _ v). - - Coercion natToBlist : nat >-> list. - - Definition lookupInds(db : DB)(w : Keyword) := - fold_left (fun ls p => if (in_dec (EqDec_dec _) w (snd p)) then (fst p :: ls) else ls) db nil. - - Local Open Scope group_scope. - - Definition OXT_EDBSetup_indLoop k_I k_Z k_E k_X w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - e <-$ Enc k_E ind; - xtag <- g^((F_P k_X w) * xind); - ret (Vector.append e (natToBvector y), xtag). - - Definition OXT_EDBSetup_wLoop db k_S k_I k_Z k_X w := - k_E <- F k_S w; - inds <- lookupInds db w; - inds <-$ shuffle _ inds; - indLoopRes <-$ compMap _ (OXT_EDBSetup_indLoop k_I k_Z k_E k_X w) (combine inds (allNatsLt (length inds))); - ret (split indLoopRes). - - Definition toW (db : DB) := - removeDups _ (flatten (snd (split db))). - - Definition OXT_EDBSetup (db : DB) : Comp (Key * (TSet * XSet)) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (OXT_EDBSetup_wLoop db k_S k_I k_Z k_X) (toW db); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - ret ((k_S, k_X, k_I, k_Z, k_T), (tSet, xSet)). - - Definition OXT_Search_ClientInit(k : Key)(q : Query) := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [w1, _] <-2 q; - TSetGetTag k_T w1. - - Definition OXT_Search_ServerInit(edb : EDB) stag := - [tSet, _] <-2 edb; - TSetRetrieve tSet stag. - - Definition computeXToken c k_Z k_X s x := - g^((F_P k_Z (s ++ c)) * (F_P k_X x)). - - Definition OXT_Search_Loop_Client (k : Key)(q : Query) c := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [s, x] <-2 q; - computeXToken c k_Z k_X s x. - - Fixpoint splitVector(A : Set)(n m : nat) : Vector.t A (n + m) -> (Vector.t A n * Vector.t A m) := - match n with - | 0 => - fun (v : Vector.t A (O + m)) => (@Vector.nil A, v) - | S n' => - fun (v : Vector.t A (S n' + m)) => - let (v1, v2) := splitVector _ _ (Vector.tl v) in - (Vector.cons _ (Vector.hd v) _ v1, v2) - end. - - Definition OXT_Search_Loop_Server (edb : EDB) xtoken t_c := - [_, xSet] <-2 edb; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition OXT_Search_Loop edb k q (e : nat * Bvector (lambda + lambda)) := - [c, t_c] <-2 e; - xtoken <- OXT_Search_Loop_Client k q c ; - answer <- OXT_Search_Loop_Server edb xtoken t_c; - (xtoken, answer). - - Definition OXT_Search_ClientFinalize (k : Key) (q : Query) es := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [w1, ws] <-2 q; - k_E <- F k_S w1; - map (Dec k_E) es. - - Fixpoint getSomes(A : Type)(ls : list (option A)) := - match ls with - | nil => nil - | o :: ls' => - match o with - | None => (getSomes ls') - | Some a => a :: (getSomes ls') - end - end. - - - Definition SearchTranscript := (Tag * list (nat * option (Bvector lambda)))%type. - - Definition OXT_Search (edb : EDB) (k : Key) (q : Query) : Comp (list (Identifier lambda) * SearchTranscript) := - stag <-$ OXT_Search_ClientInit k q; - t <- OXT_Search_ServerInit edb stag; - xscript <- map (OXT_Search_Loop edb k q) (combine (allNatsLt (length t)) t); - es <- getSomes (snd (split xscript)); - inds <- OXT_Search_ClientFinalize k q es; - ret (inds, (stag, xscript)). - - Definition dbSize(db : DB) := - fold_left (fun n e => n + length (snd e))%nat db 0%nat. - - Require Import RndListElem. - Check firstIndexOf. - - Definition equalityPattern(ls : list Keyword) := - map (fun x => firstIndexOf (EqDec_dec _) ls x 0) ls. - - Definition sizePattern(db : DB)(ls : list Keyword) := - map (fun x => length (lookupInds db x)) ls. - - Definition lookupIndsConj(db : DB)(w : Query) := - [s, x] <-2 w; - inds <- lookupInds db s; - intersect (EqDec_dec _) inds (lookupInds db x). - - Definition resultsPattern(db : DB)(q: list Query) := - map (lookupIndsConj db) q. - - Definition condIntPattern(db : DB) (i j : nat) (qi qj : Query):= - r <- intersect (EqDec_dec _) (lookupInds db (fst qi)) (lookupInds db (fst qj)); - if (eq_nat_dec i j) then nil else - if (eqb (snd qi) (snd (qj))) then r else nil. - - Definition condIntPatternTable(db : DB)(q : list Query) := - map (fun e => [i,qi] <-2 e; - map - (fun e => [j,qj] <-2 e; condIntPattern db i j qi qj) - (combine (allNatsLt (length q)) q)) - (combine (allNatsLt (length q)) q). - - Definition L_OXT(db : DB)(q : list Query) := - n <- dbSize db; - s_bar <- equalityPattern (fst (split q)); - sP <- sizePattern db (fst (split q)); - rP <- resultsPattern db q; - iP <- condIntPatternTable db q; - (n, s_bar, sP, rP, iP). - - Require Import RndGrpElem. - - Definition L_cLoop (k : Bvector lambda)(c : nat) := - y <-$ RndG; - e <-$ Enc k (Vector.const false lambda); - ret (Vector.append (natToBvector y) e). - - Definition L_wLoop (db : DB)(w : Keyword) := - k <-$ {0, 1}^lambda; - t_entries <-$ compMap _ (L_cLoop k) (allNatsLt (length (lookupInds db w))); - ret (w, t_entries). - - Definition L (db : DB)(q : list Query) := - bigT <-$ compMap _ (L_wLoop db) (toW db); - s <- fst (unzip q); - tSetAnswers <- map (lookupAnswers (pos (lambda + lambda))%nat bigT) s; - ret (L_OXT db q, L_T bigT s, tSetAnswers ). - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - (* Step 1: simplify and put the game into the "Initialize" form of the paper. *) - Definition Initialize_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (OXT_EDBSetup_wLoop db k_S k_I k_Z k_X) (toW db); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Definition G_gen (init : DB -> list Query -> Comp (EDB * (list (list (Bvector lambda) * SearchTranscript)))) := - [db, q, s_A] <-$3 A1; - [edb, ls] <-$2 init db q; - A2 s_A edb (snd (split ls)). - - Theorem G1_equiv : - Pr[SSE_Sec_NA_Real OXT_EDBSetup _ OXT_Search A1 A2 ] == Pr[G_gen Initialize_1]. - - intuition. - unfold SSE_Sec_NA_Real, G_gen. - comp_skip. - comp_simp. - unfold Initialize_1, OXT_EDBSetup. - do 7 (inline_first; comp_skip; comp_simp); - reflexivity. - - Qed. - - - (* Step 2: simplify and split some loops. This gets us close to game G0 in the paper, except we still -look up answers in the real way. *) - - Definition Initialize_indLoop_2 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y)). - - Definition Initialize_wLoop_2 db k_S k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <- F k_S w; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition XSetSetup_indLoop_2 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - g^(e * xind). - - Definition XSetSetup_wLoop_2 k_X k_I db e := - [w, sigma] <-2 e; - map (XSetSetup_indLoop_2 k_X k_I w) (permute (lookupInds db w) sigma). - - Definition XSetSetup_2 k_X k_I db sigmas := - flatten (map (XSetSetup_wLoop_2 k_X k_I db) (combine (toW db) sigmas)). - - Definition ServerSearchLoop_2 xSet e := - [xtoken, t_c] <-2 e; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition ServerSearch_2 xSet (xtokens : list nat) t := - map (ServerSearchLoop_2 xSet) (combine xtokens t). - - Definition GenTrans_2 (edb : EDB) k_X k_Z k_S (e : Query * Tag) : (list (Identifier lambda) * SearchTranscript) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_2 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - Definition G2 := G_gen Initialize_2. - - -(* Step 1.1: start splitting off the XSet computation. This will take a few steps and we'll start with the innermost loop. *) - - Definition Initialize_wLoop_1_1 db k_S k_I k_Z k_X w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute inds sigma; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - xSet <- map (XSetSetup_indLoop_2 k_X k_I w) inds; - ret (indLoopRes, xSet). - - Definition Initialize_1_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_1 db k_S k_I k_Z k_X) (toW db); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem fst_split_eq_list_pred : - forall (A B : Set)(ls1 : list (A * B))(ls2 : list A), - list_pred (fun a b => fst a = b) ls1 ls2 -> - fst (split ls1) = ls2. - - induction 1; intuition; simpl in *. - subst. - destruct a1. - remember (split ls1) as x. - destruct x. - simpl in *. - f_equal. - Qed. - - Theorem G1_1_equiv : - Pr[G_gen Initialize_1] == Pr[G_gen Initialize_1_1]. - - unfold G_gen, Initialize_1, Initialize_1_1. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_1_1_spec : - forall d x x1 x2 x0 b0, - comp_spec eq - (OXT_EDBSetup_wLoop d x x1 x2 x0 b0) - (Initialize_wLoop_1_1 d x x1 x2 x0 b0). - - intuition. - unfold OXT_EDBSetup_wLoop, Initialize_wLoop_1_1. - comp_simp. - - assert (comp_spec eq - (inds <-$ shuffle (Bvector_EqDec lambda) (lookupInds d b0); - indLoopRes <-$ - (foreach (x3 in combine inds (allNatsLt (length inds))) - OXT_EDBSetup_indLoop x1 x2 (F x b0) x0 b0 x3); - ret split indLoopRes) - (inds <-$ ( - x <-$ RndPerm (length (lookupInds d b0)); ret permute (lookupInds d b0) x - ); - indLoopRes <-$ - (foreach (x3 in combine inds (allNatsLt (length inds))) - OXT_EDBSetup_indLoop x1 x2 (F x b0) x0 b0 x3); - ret split indLoopRes) - ). - - comp_skip. - eapply shuffle_RndPerm_spec_eq. - subst. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_trans. - eapply H. - clear H. - - inline_first. - comp_skip. - - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_indLoop_2_spec : - forall x1 x2 x3 x0 b0 e, - comp_spec (fun p1 p2 => fst p1 = p2) - (OXT_EDBSetup_indLoop x1 x2 x3 x0 b0 e) - (Initialize_indLoop_2 x1 x2 x3 b0 e). - - intuition. - unfold OXT_EDBSetup_indLoop, Initialize_indLoop_2. - comp_simp. - comp_skip; try eapply (oneVector (lambda + lambda)). - eapply comp_spec_ret; intuition. - Qed. - - eapply Init_indLoop_2_spec. - eapply comp_spec_ret; intuition. - - remember (split a) as z. - destruct z. - f_equal. - assert (fst (split a) = b1). - eapply fst_split_eq_list_pred . - eauto. - rewrite <- Heqz in H4. - simpl in H4. - trivial. - - eapply (@compMap_support _ _ (fun z1 z2 => @snd (Bvector (lambda + lambda)) nat z2 = XSetSetup_indLoop_2 x0 x1 b0 (fst z1))) in H1. - - Theorem list_pred_fst_split_irr_gen : - forall (A B C : Set) (P : A -> C -> Prop) ls1 ls2 , - list_pred (fun (x : A * B)(y : C) => P (fst x) y) ls1 ls2 -> - forall ls3, ls3 = fst (split ls1) -> - list_pred (fun (x : A)(y : C) => P x y) ls3 ls2. - - induction 1; intuition; simpl in *; subst. - econstructor. - destruct a1. - remember (split ls1) as z. - destruct z. - simpl in *. - econstructor. - eauto. - eapply IHlist_pred. - intuition. - Qed. - - Theorem list_pred_fst_split_irr : - forall (A B C : Set) (P : A -> C -> Prop) ls1 ls2 , - list_pred (fun (x : A * B)(y : C) => P (fst x) y) ls1 ls2 -> - list_pred (fun (x : A)(y : C) => P x y) (fst (split ls1)) ls2. - - intuition. - eapply list_pred_fst_split_irr_gen; eauto. - - Qed. - - Theorem list_pred_snd_split_irr_gen : - forall (A B C : Set) (P : B -> C -> Prop) ls1 ls2 , - list_pred (fun (x : A * B)(y : C) => P (snd x) y) ls1 ls2 -> - forall ls3, ls3 = snd (split ls1) -> - list_pred (fun (x : B)(y : C) => P x y) ls3 ls2. - - induction 1; intuition; simpl in *; subst. - econstructor. - destruct a1. - remember (split ls1) as z. - destruct z. - simpl in *. - econstructor. - eauto. - eapply IHlist_pred. - intuition. - Qed. - - - Theorem list_pred_snd_split_irr : - forall (A B C : Set) (P : B -> C -> Prop) ls1 ls2 , - list_pred (fun (x : A * B)(y : C) => P (snd x) y) ls1 ls2 -> - list_pred (fun (x : B)(y : C) => P x y) (snd (split ls1)) ls2. - - intuition. - eapply list_pred_snd_split_irr_gen; eauto. - Qed. - - Theorem snd_split_eq_map_list_pred : - forall (A B C : Set)(ls1 : list (A * C))(ls2 : list B)(f : B -> C), - list_pred (fun a b => snd a = (f b)) ls1 ls2 -> - snd (split ls1) = (map f ls2). - - induction 1; intuition; simpl in *. - destruct a1. - remember (split ls1) as x. - destruct x. - simpl in *. - subst. - f_equal. - Qed. - - Show. - assert (snd (split a) = (foreach (x in permute (lookupInds d b0) b) - XSetSetup_indLoop_2 x0 x1 b0 x)). - eapply snd_split_eq_map_list_pred. - eapply list_pred_symm. - eapply list_pred_fst_split_irr_gen. - eapply list_pred_impl. - eapply H1. - intuition. - rewrite combine_split. - simpl. - trivial. - rewrite allNatsLt_length. - trivial. - rewrite <- Heqz in H4. - simpl in *. - intuition. - - intuition. - unfold OXT_EDBSetup_indLoop, XSetSetup_indLoop_2 in *. - repeat simp_in_support. - simpl. - trivial. - Qed. - - eapply Init_wLoop_1_1_spec. - intuition. - apply list_pred_eq_impl_eq in H6. - subst. - eapply comp_spec_eq_refl. - Qed. - - (* Step 1.2: pull the XSet computation outside of the setup loop. Start by putting it in the form where we can use loop fission. *) - Definition Initialize_wLoop_1_2 db k_S k_I k_Z w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute inds sigma; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - ret (w, indLoopRes, inds). - - Definition XSetSetup_wLoop_1_2 k_X k_I (e : Blist * list (Bvector (lambda + lambda)) * list (Bvector lambda)) := - [w, x, inds] <-3 e; - r <- map (XSetSetup_indLoop_2 k_X k_I w) inds; - (x, r). - - Definition Initialize_1_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ (x <-$ compMap _ (Initialize_wLoop_1_2 db k_S k_I k_Z) (toW db); - ret (map (XSetSetup_wLoop_1_2 k_X k_I) x)); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem compMap_map_fission - : forall (A B C : Set) (eqdb : EqDec B) - (eqdc : EqDec C) (ls : list A) (f1 : A -> Comp B) - (f2 : A -> Comp C) (f3 : B -> C), - (forall a : A, - comp_spec (fun (b : B) (c : C) => f3 b = c) (f1 a) (f2 a)) -> - comp_spec eq - (lsb <-$ compMap _ f1 ls; ret (map f3 lsb)) - (compMap _ f2 ls). - - induction ls; intuition; simpl in *. - comp_simp. - simpl. - eapply comp_spec_eq_refl. - - inline_first. - - eapply comp_spec_seq; try eapply nil. - eauto. - intuition. - subst. - inline_first. - - assert ( comp_spec eq - (a1 <-$ (foreach (x in ls)f1 x); - lsb <-$ ret a0 :: a1; ret (foreach (x in lsb)f3 x)) - (a1 <-$ (foreach (x in ls)f1 x); - ret (foreach (x in (a0 :: a1))f3 x)) - ). - comp_skip. - eapply comp_spec_eq_trans. - eapply H2. - clear H2. - simpl. - - assert ( comp_spec eq - (ls <-$ (a1 <-$ (foreach (x in ls)f1 x); ret (map f3 a1)); ret f3 a0 :: ls) - (lsb' <-$ (foreach (x in ls)f2 x); ret f3 a0 :: lsb')). - - eapply comp_spec_seq_eq; try eapply nil. - eapply IHls. - intuition. - intuition. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply H2. - clear H2. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - - Theorem G_1_2_equiv : - Pr[G_gen Initialize_1_1] == Pr[G_gen Initialize_1_2]. - - unfold G_gen, Initialize_1_1, Initialize_1_2. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq_eq; eauto with inhabited. - - eapply comp_spec_eq_symm. - eapply compMap_map_fission. - intuition. - unfold Initialize_wLoop_1_1, Initialize_wLoop_1_2, XSetSetup_wLoop_1_2. - comp_simp. - comp_skip. - comp_skip. - eapply comp_spec_ret; intuition. - - intuition. - eapply comp_spec_eq_refl. - Qed. - - (* Step 1.3: Now simplify and take the game out of the form of the loop fission argument*) - Definition Initialize_wLoop_1_3 db k_S k_I k_Z w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute inds sigma; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - ret (indLoopRes, inds). - - Definition XSetSetup_wLoop_1_3 k_X k_I e := - [w, inds] <-2 e; - map (XSetSetup_indLoop_2 k_X k_I w) inds. - - Definition Initialize_1_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_3 db k_S k_I k_Z) (toW db); - [t_entries, inds] <-2 split wLoopRes; - xSet_raw <- map (XSetSetup_wLoop_1_3 k_X k_I) (combine (toW db) inds); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem G_1_3_equiv : - Pr[G_gen Initialize_1_2] == Pr[G_gen Initialize_1_3]. - - unfold G_gen, Initialize_1_2, Initialize_1_3. - do 6 (comp_skip; comp_simp). - inline_first. - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_equiv_1_3 : - forall d x x1 x2 b0, - comp_spec (fun a1 a2 => snd (fst a1) = fst a2 /\ snd a1 = snd a2) - (Initialize_wLoop_1_2 d x x1 x2 b0) - (Initialize_wLoop_1_3 d x x1 x2 b0). - - intuition. - - unfold Initialize_wLoop_1_2, Initialize_wLoop_1_3. - comp_simp. - comp_skip. - comp_skip. - eapply comp_spec_ret; intuition. - - Qed. - - eapply Init_wLoop_equiv_1_3. - intuition. - comp_ret_l. - remember (split (map (XSetSetup_wLoop_1_2 x0 x1) a1)) as z. - destruct z. - remember (split b0) as z. - destruct z. - - - Theorem XSetSetup_wLoop_1_2_fst_eq : - forall a1 x0 x1, - fst (split (map (XSetSetup_wLoop_1_2 x0 x1) a1)) = snd (split (fst (split a1))). - - induction a1; intuition; simpl in *. - remember (split (foreach (x in a1)XSetSetup_wLoop_1_2 x0 x1 x))as z. - destruct z. - remember (split a1) as z. - destruct z. - simpl. - remember (split l1) as z. - destruct z. - simpl. - f_equal. - assert (l = (fst (split (foreach (x in a1)XSetSetup_wLoop_1_2 x0 x1 x)))). - rewrite <- Heqz. - trivial. - subst. - rewrite IHa1. - simpl. - rewrite <- Heqz1. - trivial. - Qed. - - Show. - assert (l0 = l2). - assert (l0 = fst (split (map (XSetSetup_wLoop_1_2 x0 x1) a1))). - rewrite <- Heqz. - trivial. - subst. - rewrite XSetSetup_wLoop_1_2_fst_eq. - - assert (fst (split b0) = l2). - rewrite <- Heqz0. - trivial. - subst. - - - Ltac split_eq_one := - match goal with - | [|- _ = _] => eapply list_pred_eq_impl_eq - | [|- list_pred _ (fst (split _)) _ ] => eapply list_pred_fst_split_irr - | [|- list_pred _ (snd (split _)) _ ] => eapply list_pred_snd_split_irr - | [|- list_pred _ _ (fst (split _)) ] => eapply list_pred_symm; split_eq_one; eapply list_pred_symm - | [|- list_pred _ _ (snd (split _)) ] => eapply list_pred_symm; split_eq_one; eapply list_pred_symm - end. - - Ltac split_eq := repeat split_eq_one. - - split_eq. - eapply list_pred_impl. - eauto. - intuition. - - subst. - comp_skip. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem comp_spec_eq_refl_gen : - forall (A : Set)(eqd : EqDec A)(c1 c2 : Comp A), - c1 = c2 -> - comp_spec eq c1 c2. - - intuition; subst. - eapply comp_spec_eq_refl. - - Qed. - - eapply comp_spec_eq_refl_gen. - f_equal. - f_equal. - f_equal. - - Ltac split_subst := - match goal with - | [H : (_, ?b) = split ?ls |- context[?b] ] => assert (b = (snd (split ls))); [rewrite <- H; trivial | idtac]; subst - | [H : (?a, _) = split ?ls |- context[?a] ] => assert (a = (fst (split ls))); [rewrite <- H; trivial | idtac]; subst - end. - - split_subst. - split_eq. - - eapply (@list_pred_impl _ _ _ _ (fun a b => snd a = b)). - - Theorem list_pred_map_both': - forall (A B C D: Set) (lsa : list A) (lsb : list B) - (P : C -> D -> Prop) (f1 : A -> C)(f2 : B -> D), - list_pred (fun (a : A) (b : B) => P (f1 a) (f2 b)) lsa lsb -> - list_pred P (map f1 lsa) (map f2 lsb). - - induction 1; intuition; simpl in *. - econstructor. - econstructor; eauto. - - Qed. - - eapply (list_pred_map_both'). - - Theorem list_pred_combine_l_h : - forall (A C : Set)(lsa : list A)(lsc : list C) P1, - list_pred P1 lsa lsc -> - forall (B : Set)(lsb : list B) P2, - list_pred P2 lsb lsc -> - list_pred (fun p c => P1 (fst p) c /\ P2 (snd p) c) (combine lsa lsb) lsc. - - induction 1; intuition; simpl in *. - econstructor. - inversion H1; clear H1; subst. - econstructor. - intuition. - - eauto. - - Qed. - - Theorem list_pred_combine_l : - forall (A B C : Set)P1 P2 (lsa : list A)(lsb : list B)(lsc : list C), - list_pred P1 lsa lsc -> - list_pred P2 lsb lsc -> - list_pred (fun p c => P1 (fst p) c /\ P2 (snd p) c) (combine lsa lsb) lsc. - - intuition. - eapply list_pred_combine_l_h; eauto. - - Qed. - - - eapply list_pred_symm. - eapply (@list_pred_impl). - - eapply (@list_pred_combine_l _ _ _ _ (fun a b => a = snd b)). - - eapply (compMap_support (fun a b => fst (fst b) = a)). - eapply H4. - intuition. - unfold Initialize_wLoop_1_2 in *. - repeat simp_in_support. - trivial. - - split_subst. - eapply (@list_pred_impl _ _ _ _ (fun a b => a = snd b)). - split_eq. - eapply list_pred_symm. - eapply (@list_pred_impl). - eapply H6. - intuition. - intuition. - - intuition. - simpl in *; subst. - unfold XSetSetup_wLoop_1_2. - comp_simp. - simpl. - trivial. - - intuition. - eapply comp_spec_ret; intuition. - f_equal. - f_equal. - f_equal. - split_subst. - split_eq. - eapply list_pred_map_both'. - eapply list_pred_symm. - eapply list_pred_impl. - eapply (@list_pred_combine_l _ _ _ _ (fun a b => a = snd b)). - eapply (compMap_support (fun a b => fst (fst b) = a)). - eapply H4. - intuition. - unfold Initialize_wLoop_1_2 in *. - repeat simp_in_support. - trivial. - - split_subst. - eapply (@list_pred_impl _ _ _ _ (fun a b => a = snd b)). - split_eq. - eapply list_pred_symm. - eapply (@list_pred_impl). - eapply H6. - intuition. - intuition. - intuition. - simpl in *; subst. - unfold XSetSetup_wLoop_1_2. - comp_simp. - trivial. - - eapply list_pred_eq_impl_eq. - trivial. - - Qed. - - (* Step 1.4: have the setup routine return permutations instead of indices. Start by returning both. *) - Definition Initialize_wLoop_1_4 db k_S k_I k_Z w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute inds sigma; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - ret (indLoopRes, inds, sigma). - - Definition Initialize_1_4 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_4 db k_S k_I k_Z) (toW db); - [t_inds, sigmas] <-2 split wLoopRes; - [t_entries, inds] <-2 split t_inds; - xSet_raw <- map (XSetSetup_wLoop_1_3 k_X k_I) (combine (toW db) inds); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem G_1_4_equiv : - Pr[G_gen Initialize_1_3] == Pr[G_gen Initialize_1_4]. - - unfold G_gen, Initialize_1_3, Initialize_1_4. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_1_4_spec : - forall d x x1 x2 b0, - comp_spec (fun a b => fst a = fst (fst b) /\ snd a = snd (fst b)) - (Initialize_wLoop_1_3 d x x1 x2 b0) - (Initialize_wLoop_1_4 d x x1 x2 b0). - - intuition. - unfold Initialize_wLoop_1_3, Initialize_wLoop_1_4. - comp_simp. - comp_skip. - comp_skip. - eapply comp_spec_ret; intuition. - - Qed. - - eapply Init_wLoop_1_4_spec. - remember (split a1) as z. - destruct z. - remember ( split b0)as z. - destruct z. - remember ( split l2) as z. - destruct z. - assert (l0 = l4). - do 3 (split_subst; split_eq). - eapply list_pred_impl. - eauto. - intuition. - subst. - comp_skip. - assert (l1 = l5). - do 3 (split_subst; split_eq). - eapply list_pred_impl. - eauto. - intuition. - subst. - - intuition. - - Qed. - - (* Step 1.5: Now remove the indices and just use the permutations *) - Definition Initialize_1_5 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet_raw <- map (XSetSetup_wLoop_2 k_X k_I db) (combine (toW db) sigmas); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem G_1_5_equiv : - Pr[G_gen Initialize_1_4] == Pr[G_gen Initialize_1_5]. - - unfold G_gen, Initialize_1_4, Initialize_1_5. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_2_spec : - forall d x x1 x2 b0, - comp_spec (fun a b => fst (fst a) = fst b /\ snd a = snd b) - (Initialize_wLoop_1_4 d x x1 x2 b0) - (Initialize_wLoop_2 d x x1 x2 b0). - - intuition. - unfold Initialize_wLoop_1_4, Initialize_wLoop_2. - comp_simp. - comp_skip. - comp_skip. - rewrite permute_length_eq. - apply RndPerm_In_support_length in H0. - rewrite H0. - eapply comp_spec_eq_refl. - intuition. - eapply allNatsLt_lt. - - Require Import Permutation. - - eapply Permutation_in. - eapply Permutation_sym. - eapply RndPerm_In_support. - eauto. - trivial. - subst. - - eapply comp_spec_ret; intuition. - - Qed. - - eapply Init_wLoop_2_spec. - remember (split a1) as z. - destruct z. - remember (split l0) as z. - destruct z. - remember (split b0) as z. - destruct z. - - assert (l2 = l4). - do 3 (split_subst; split_eq). - eapply list_pred_impl. - eauto. - intuition. - subst. - comp_skip. - - assert (l1 = l5). - do 2 (split_subst; split_eq). - eapply list_pred_impl. - eauto. - intuition. - subst. - - assert (map (XSetSetup_wLoop_1_3 x0 x1) (combine (toW d) l3) = - (map (XSetSetup_wLoop_2 x0 x1 d) (combine (toW d) l5))). - - eapply (compMap_support (fun a b => snd (fst b) = permute (lookupInds d a) (snd b))) in H4. - - Theorem map_ext_pred : - forall (A B C : Set)(P : A -> B -> Prop)(lsa : list A)(lsb : list B)(f1 : A -> C)(f2 : B -> C), - list_pred P lsa lsb -> - (forall a b, P a b -> (f1 a) = (f2 b)) -> - map f1 lsa = map f2 lsb. - - induction 1; intuition; simpl in *. - f_equal; intuition. - - Qed. - - eapply (@map_ext_pred _ _ _ (fun a b => fst a = fst b /\ snd a = permute (lookupInds d (fst a)) (snd b))). - - - Theorem list_pred_combine_same : - forall (A B C : Set)(P : A -> (B * C) -> Prop)(lsa : list A)(lsb : list B)(lsc : list C), - list_pred P lsa (combine lsb lsc) -> - length lsb = length lsc -> - list_pred (fun p1 p2 => fst p1 = fst p2 /\ P (fst p1) (snd p1, snd p2)) - (combine lsa lsb) - (combine lsa lsc). - - induction lsa; intuition; simpl in *. - econstructor. - - inversion H; subst; clear H. - - destruct lsb; destruct lsc; simpl in *; try congruence. - inversion H3; clear H3; subst. - econstructor. - - simpl; intuition. - eauto. - - Qed. - - - eapply list_pred_impl. - eapply (@list_pred_combine_same _ _ _ (fun a2 b2 => (fst b2) = permute (lookupInds d a2) (snd b2))). - assert (combine l3 l5 = map (fun z => (snd (fst z), snd z)) a1). - - Theorem combine_eq_map_1_3 : - forall (A B C : Set)(a : list (A * B * C))(l1 : list B)(l2 : list C), - l1 = snd (split (fst (split a))) -> - l2 = snd (split a) -> - combine l1 l2 = map (fun p => (snd (fst p), snd p)) a. - - induction a; intuition; subst; simpl in *. - trivial. - - remember (split a0 )as z. - destruct z. - simpl. - remember (split l) as z. - destruct z. - simpl. - f_equal. - eapply IHa. - - simpl. - rewrite <- Heqz0. - trivial. - trivial. - Qed. - - eapply combine_eq_map_1_3. - do 2 (split_subst; split_eq). - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst; trivial. - - (split_subst; split_eq). - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst; trivial. - - rewrite H9. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply H4. - intuition. - - repeat (split_subst; try rewrite split_length_l; try rewrite split_length_r); trivial. - - intuition. - intuition. - simpl in *; subst. - intuition. - - intuition. - unfold Initialize_wLoop_1_4 in *. - repeat simp_in_support. - trivial. - - rewrite H9. - eapply comp_spec_eq_refl. - Qed. - - - (* Step 1.6: Now we factor out the tag computation. I did this without using the loop fission argument for some reason. *) - - Definition Initialize_1_6 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet_raw <- map (XSetSetup_wLoop_2 k_X k_I db) (combine (toW db) sigmas); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ - (ls <-$ compMap _ (fun p => z <-$ TSetGetTag k_T (fst p); ret (p, z)) q; - ret (map (GenTrans_2 edb k_X k_Z k_S) ls)); - ret (edb, searchRes). - - Theorem G_1_6_equiv : - Pr[G_gen Initialize_1_5] == Pr[G_gen Initialize_1_6]. - - unfold G_gen, Initialize_1_5, Initialize_1_6. - do 9 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_eq_symm. - eapply compMap_map_fission. - intuition. - - unfold OXT_Search. - unfold OXT_Search_ClientInit. - comp_simp. - simpl. - - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - - eapply comp_spec_ret; intuition. - unfold GenTrans_2. - comp_simp. - unfold ServerSearch_2. - f_equal. - f_equal. - f_equal. - - split_eq. - eapply list_pred_map_both'. - - Theorem list_pred_combine_simple : - forall (A B C D : Set)(P1 : A -> C -> Prop)(P2 : B -> D -> Prop) lsa lsb lsc lsd, - list_pred P1 lsa lsc -> - list_pred P2 lsb lsd -> - length lsa = length lsb -> - list_pred (fun x1 x2 => P1 (fst x1) (fst x2) /\ P2 (snd x1) (snd x2)) - (combine lsa lsb) - (combine lsc lsd). - - induction lsa; intuition; simpl in *. - inversion H; clear H; subst. - simpl. - econstructor. - - destruct lsb; simpl in *; try discriminate. - - inversion H; clear H; subst. - inversion H0; clear H0; subst. - simpl. - econstructor; simpl in *; eauto. - - Qed. - - Show. - eapply list_pred_impl. - eapply (@list_pred_combine_simple _ _ _ _ (fun (a c : nat) => a = g ^ (F_P x0 k0 * F_P x2 (k ++ c))) eq). - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst; intuition. - eapply list_pred_eq. - rewrite map_length. - eapply allNatsLt_length. - - intuition. - unfold ServerSearchLoop_2, OXT_Search_Loop. - unfold OXT_Search_Loop_Client, OXT_Search_Loop_Server, computeXToken. - simpl in *; subst. - destruct b2. - simpl. - comp_simp. - rewrite mult_comm. - trivial. - - f_equal. - - Theorem combine_map_eq : - forall (A B C D : Type)(f1 : A -> B)(f2 : C -> D) ls1 ls2, - (length ls1 = length ls2) -> - combine (map f1 ls1) (map f2 ls2) = map (fun p => (f1 (fst p), f2 (snd p))) (combine ls1 ls2). - - induction ls1; destruct ls2; intuition; simpl in *. - f_equal; eauto. - - Qed. - - rewrite combine_map_eq. - - eapply (@map_ext_pred _ _ _ (fun (a : nat * (nat * Bvector (lambda + lambda))) b => fst a = fst b /\ snd (snd a) = snd b - /\ fst (snd a) = (g ^ (F_P x2 (k ++ (fst b)) * F_P x0 k0)) )). - - Theorem list_pred_combine_same' - : forall (A B C : Set) (P : A * B -> A * C -> Prop) - (lsa : list A) (lsb : list B) (lsc : list C), - list_pred (fun a p => P (a, fst p) (a, snd p)) lsa (combine lsb lsc) -> - length lsb = length lsc -> - list_pred - P - (combine lsa lsb) (combine lsa lsc). - - induction lsa; intuition; simpl in *. - econstructor. - - destruct lsb; destruct lsc; simpl in *; try discriminate. - econstructor. - - inversion H; clear H; subst. - econstructor; eauto. - - Qed. - - intuition. - eapply list_pred_combine_same'. - eapply list_pred_symm. - - - Theorem list_pred_combine_assoc_l : - forall (A B C D : Set)(P : (A * B) * C -> D -> Prop) lsa lsb lsc lsd, - list_pred (fun p d => P ((fst p, fst (snd p)), snd (snd p)) d) - (combine lsa (combine lsb lsc)) - lsd -> - list_pred P - (combine (combine lsa lsb) lsc) - lsd. - - induction lsa; intuition; simpl in *. - inversion H; clear H; subst. - econstructor. - - destruct lsb; simpl in *. - inversion H; clear H; subst. - econstructor. - - destruct lsc; simpl in *. - inversion H; clear H; subst. - econstructor. - - inversion H; clear H; subst. - econstructor. - simpl in *; trivial. - - eauto. - - Qed. - - eapply list_pred_combine_assoc_l. - - assert (list_pred (fun (b : (nat * (Bvector (lambda + lambda) * (Bvector (lambda + lambda))))) (a : nat) => fst b = (g ^ (F_P x0 k0 * F_P x2 (k ++ a))) /\ fst (snd b) = snd (snd b)) - (combine - (for (c '< length (TSetRetrieve t b0)) - g ^ (F_P x0 k0 * F_P x2 (k ++ (natToBlist c)))) - (combine (TSetRetrieve t b0) (TSetRetrieve t b0))) - (allNatsLt (length (TSetRetrieve t b0))) - ). - - eapply list_pred_impl. - - eapply (@list_pred_combine_l _ _ _ (fun (a : nat) (b : nat) => a = g ^ (F_P x0 k0 * F_P x2 (k ++ b))) (fun a b => fst a = snd a)); intuition. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst; intuition. - - Theorem list_pred_combine_same_l : - forall (A B : Set)(P : A -> B -> Prop) lsa lsb, - list_pred P lsa lsb -> - list_pred (fun p b => fst p = snd p /\ (P (fst p) b)) (combine lsa lsa) lsb. - - induction lsa; intuition; simpl in *. - inversion H; clear H; subst. - econstructor. - - inversion H; clear H; subst. - econstructor; eauto. - - Qed. - - eapply list_pred_impl. - eapply list_pred_combine_same_l. - eapply list_pred_I. - symmetry. - eapply allNatsLt_length. - intuition. - intuition. - - eapply list_pred_impl. - eapply H8. - intuition. - simpl in *. - subst. - rewrite mult_comm. - trivial. - - rewrite combine_length. - rewrite Min.min_l. - rewrite map_length. - eapply allNatsLt_length. - rewrite map_length. - rewrite allNatsLt_length. - intuition. - - intuition. - simpl in *; subst; intuition. - intuition. - unfold ServerSearchLoop_2, OXT_Search_Loop, OXT_Search_Loop_Client, computeXToken. - simpl. - comp_simp. - f_equal. - rewrite mult_comm. - trivial. - - rewrite combine_length. - rewrite Min.min_l. - rewrite map_length. - trivial. - rewrite map_length. - rewrite allNatsLt_length. - intuition. - Qed. - - (* We get G2 by simplification *) - Theorem G2_equiv : - Pr[G_gen Initialize_1_6] == Pr[G2]. - - unfold G2, G_gen, Initialize_1_6, Initialize_2. - do 8 (comp_skip; comp_simp). - inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply (@compMap_spec _ _ _ _ _ _ (fun a b => fst a = b) (fun a b => snd a = b)). - split_eq. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst. - trivial. - intuition. - subst. - eapply comp_spec_symm. - eapply comp_spec_eq_trans_l. - eapply comp_spec_symm. - eapply comp_spec_consequence. - eapply comp_spec_right_ident. - intuition. - eapply comp_spec_seq_eq. - eapply comp_base_exists; eauto. - intuition. - eapply comp_base_exists; eauto. - eapply comp_spec_eq_refl. - intuition. - eapply comp_spec_ret; intuition. - - eapply comp_spec_ret; intuition. - f_equal. - - eapply (@compMap_support _ _ (fun a b => (fst b) = a)) in H6. - - eapply (@map_ext_pred _ _ _ eq). - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply H6. - eapply list_pred_symm. - eapply H8. - intuition. - simpl in *; subst. - destruct b2; trivial. - - intuition; subst. - trivial. - - intuition. - repeat simp_in_support. - trivial. - - Qed. - - (* Step 3 : replace the TSetRetrieve and use the actual values instead. We make this transformation by applying the TSet correctness definition. *) - Definition GenTrans_3 (edb : EDB) k_X k_Z k_S (e : Query * (Tag * list (Bvector (lambda + lambda)))) : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_3 edb k_X k_Z k_S) (combine q stags_ts); - ret (edb, searchRes). - - Definition G3 := G_gen Initialize_3. - - (*Step 2.1: factor the TSetRetrieve out into the main loop. *) - Definition Initialize_2_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - ts <- map (TSetRetrieve tSet) stags; - stags_ts <- combine stags ts; - searchRes <- map (GenTrans_3 edb k_X k_Z k_S) (combine q stags_ts); - ret (edb, searchRes). - - - Theorem G2_1_equiv : - Pr[G2] == Pr[G_gen Initialize_2_1]. - - unfold G2, G_gen, Initialize_2, Initialize_2_1. - do 9 (comp_skip; comp_simp). - eapply evalDist_ret_eq. - f_equal. - eapply (@map_ext_pred). - - eapply (@list_pred_combine_simple _ _ _ _ eq (fun b a => fst a = b /\ snd a = TSetRetrieve t b)). - eapply list_pred_eq. - eapply list_pred_symm. - eapply (@list_pred_impl _ _ _ _ (fun a b => fst a = b /\ snd a = TSetRetrieve t b)). - eapply (@list_pred_combine_l _ _ _ eq (fun a b => a = TSetRetrieve t b)). - eapply list_pred_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst; trivial. - - intuition. - apply compMap_length in H6. - rewrite H6. - symmetry. - eapply split_length_l. - - intuition. - simpl in *; subst. - unfold GenTrans_3. - destruct b1. - destruct p0. - simpl in *. - subst. - trivial. - - Qed. - - - (*Step 2.2: We will use an identical until bad argument. Make the badness from the top-level game. *) - Definition G2_2 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - ts <- map (TSetRetrieve tSet) stags; - correct_ts <- map (arrayLookupList _ t) (fst (split q)); - stags_ts <- combine stags ts; - searchRes <- map (GenTrans_3 edb k_X k_Z k_S) (combine q stags_ts); - b <-$ A2 s_A edb (snd (split searchRes)); - ret (b, negb (eqb ts correct_ts)). - - Theorem G2_2_equiv : - Pr[G_gen Initialize_2_1] == Pr[p <-$ G2_2; ret fst p]. - - unfold G_gen, Initialize_2_1, G2_2. - repeat (inline_first; comp_skip; comp_simp). - inline_first. - rewrite <- evalDist_right_ident. - comp_skip. - reflexivity. - comp_simp. - simpl. - intuition. - Qed. - - - (*Step 2.3: switch over to the actual values*) - Definition G2_3 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - ts <- map (TSetRetrieve tSet) stags; - correct_ts <- map (arrayLookupList _ t) (fst (split q)); - stags_ts <- combine stags correct_ts; - searchRes <- map (GenTrans_3 edb k_X k_Z k_S) (combine q stags_ts); - b <-$ A2 s_A edb (snd (split searchRes)); - ret (b, negb (eqb ts correct_ts)). - - Theorem G2_2_3_badness_eq : - Pr [a <-$ G2_2; ret snd a ] == Pr [a <-$ G2_3; ret snd a ]. - - unfold G2_2, G2_3. - do 8 (inline_first ; comp_skip; comp_simp). - inline_first. - comp_irr_l. - comp_irr_r. - - comp_simp. - simpl. - intuition. - - Qed. - - Theorem G2_2_3_eq_until_bad : forall x, - evalDist (a <-$ G2_2; ret (fst a, snd a)) (x, false) == - evalDist (a <-$ G2_3; ret (fst a, snd a)) (x, false). - - intuition. - unfold G2_2, G2_3. - - do 8 (inline_first ; comp_skip; comp_simp). - case_eq ((eqb (map (TSetRetrieve t) x5) - (map - (arrayLookupList (list_EqDec bool_EqDec) - (combine (toW d) l0)) (fst (split l))))); intuition. - - rewrite eqb_leibniz in H7. - repeat rewrite H7. - reflexivity. - - inline_first. - comp_irr_l. - comp_irr_r. - comp_simp. - repeat rewrite evalDist_ret_0; intuition. - - pairInv. - apply negb_false_iff in H13. - rewrite eqb_leibniz in H13. - rewrite H13 in H7. - rewrite eqb_refl in H7. - discriminate. - - pairInv. - apply negb_false_iff in H13. - rewrite eqb_leibniz in H13. - rewrite H13 in H7. - rewrite eqb_refl in H7. - discriminate. - Qed. - - Theorem G2_3_equiv : - | Pr[x <-$ G2_2; ret fst x] - Pr[x <-$ G2_3; ret fst x] | - <= Pr[x <-$ G2_2; ret snd x]. - - eapply fundamental_lemma; intuition. - eapply G2_2_3_badness_eq. - eapply G2_2_3_eq_until_bad. - - Qed. - - - (* Now we need to show that the bad event is related to TSet correctness. Put the game in the correct form. *) - Definition TSetCorA := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - ret (t, (fst (split q))). - - Definition G2_2_bad1 := - [t, q] <-$2 TSetCorA; - [tSet, k_T] <-$2 TSetSetup t; - stags <-$ compMap _ (TSetGetTag k_T) q; - ts <- map (TSetRetrieve tSet) stags; - correct_ts <- map (arrayLookupList _ t) q; - ret (negb (eqb ts correct_ts)). - - Theorem G2_2_bad1_equiv : - Pr[x <-$ G2_2; ret snd x] == - Pr[G2_2_bad1]. - - unfold G2_2_bad1, TSetCorA, G2_2. - do 2 (inline_first; comp_skip; comp_simp). - comp_inline_l. - comp_irr_l; wftac. - do 5 (inline_first; comp_skip; comp_simp). - - comp_inline_l. - comp_irr_l. - eapply evalDist_ret_eq; simpl; intuition. - Qed. - - Theorem G2_2_bad1_TSetCor : - Pr[G2_2_bad1] == Pr[AdvCor _ TSetSetup TSetGetTag TSetRetrieve TSetCorA]. - - unfold G2_2_bad1, TSetCorA, AdvCor. - - repeat (inline_first; comp_skip; comp_simp). - (* doesn't quite work because the definition only allows the adversary -to provide a set. *) - (* I need to do more work on this, but I may change the definition anyway when I -use the TSet later. *) - - Admitted. - - (* remove the badness and move some things around to get G3 *) - Theorem G3_equiv : - Pr[x <-$ G2_3; ret fst x] == Pr[G3]. - - unfold G2_3, G3, G_gen, Initialize_3. - (* There is some loop fission in here, but otherwise they are identical *) - - Admitted. - - - Theorem G2_G3_close : - | Pr[G2] - Pr[G3] | <= Pr[AdvCor _ TSetSetup TSetGetTag TSetRetrieve TSetCorA]. - - rewrite G2_1_equiv. - rewrite G2_2_equiv. - rewrite <- G3_equiv. - rewrite G2_3_equiv. - rewrite G2_2_bad1_equiv. - rewrite G2_2_bad1_TSetCor. - intuition. - Qed. - - (* Step 4: Give the plaintexts to GenTrans and remove the decryption.*) - Definition Initialize_indLoop_4 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_4 db k_S k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <- F k_S w; - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition ServerSearchLoop_4 xSet (e : nat * (Bvector (lambda + lambda) * Bvector lambda)) := - [xtoken, t_c_ind] <-2 e; - [t_c, ind] <-2 t_c_ind; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some (e, ind)) else None. - - Definition ServerSearch_4 xSet (xtokens : list nat) (t : list (Bvector (lambda + lambda) * Bvector lambda)) := - map (ServerSearchLoop_4 xSet) (combine xtokens t). - - Definition GenTrans_4 (edb : EDB) k_X k_Z (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_4 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_4 db k_S k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G4 := G_gen Initialize_4. - - - Theorem splitVector_append : - forall (A : Set)(n1 : nat)(v1 : Vector.t A n1)(n2 : nat)(v2 : Vector.t A n2), - splitVector n1 n2 (Vector.append v1 v2) = (v1, v2). - - induction v1; intuition; simpl in *. - - rewrite IHv1. - trivial. - Qed. - - Theorem Init_indLoop_4_spec : - forall k1 k2 k3 w id n z1 z2, - In (z1, z2) (getSupport (Initialize_indLoop_4 k1 k2 k3 w (id, n))) -> - Dec k3 (fst (splitVector lambda lambda z1)) = z2. - - intuition. - unfold Initialize_indLoop_4 in *. - repeat simp_in_support. - rewrite splitVector_append. - simpl. - eapply Enc_correct; eauto. - - Qed. - - - Theorem list_pred_impl_1_r : - forall (A B : Set)(P' : A -> B -> Prop)(P : B -> B -> Prop) ls ls', - list_pred P' ls' ls -> - (forall a b, P' a b -> P b b) -> - list_pred P ls ls. - - induction ls; intuition; simpl in *. - econstructor. - inversion H; clear H; subst. - econstructor; eauto. - - Qed. - - Theorem Init_wLoop_4_spec : - forall db k1 k2 k3 w z1 z2, - In (z1, z2) (getSupport (Initialize_wLoop_4 db k1 k2 k3 w)) -> - list_pred (fun a b => Dec (F k1 w) (fst (splitVector lambda lambda a)) = b) (fst (split z1)) (snd (split z1)). - - intuition. - unfold Initialize_wLoop_4 in *. - repeat simp_in_support. - simpl. - split_eq. - eapply list_pred_impl_1_r. - eapply (@compMap_support _ _ (fun a b => Dec (F k1 w) (fst (splitVector lambda lambda (fst b))) = (snd b))) in H1. - eauto. - intuition. - destruct b0. - simpl. - eapply Init_indLoop_4_spec. - eauto. - intuition. - Qed. - - - Definition GenTrans_3_1 (edb : EDB) k_X k_Z k_S (e : Query * (Tag * list (Bvector (lambda + lambda) * Bvector lambda))) : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - res <- map (fun x => match x with - | Some v => Some (fst v) - | None => None - end) res; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_3_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_4 db k_S k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_3_1 edb k_X k_Z k_S) (combine q stags_ts); - ret (edb, searchRes). - - Theorem list_pred_fst_split_irr_if: - forall (A B C : Set) (P : A -> C -> Prop) (ls1 : list (A * B)) - (ls2 : list C), - list_pred (fun (x : A) (y : C) => P x y) (fst (split ls1)) ls2 -> - list_pred (fun (x : A * B) (y : C) => P (fst x) y) ls1 ls2. - - induction ls1; intuition; simpl in *. - inversion H; clear H; subst. - econstructor. - remember (split ls1) as z. - destruct z. - simpl in *. - inversion H; clear H; subst. - econstructor; - simpl; - intuition. - - Qed. - - Theorem G4_equiv : - Pr[G_gen Initialize_3_1] == Pr[G4]. - - unfold G4, G_gen, Initialize_3_1, Initialize_4. - do 6 (comp_skip; comp_simp). - comp_skip. - remember (split x3) as z. - comp_simp. - comp_skip; comp_simp. - comp_skip. - - eapply (@compMap_support _ _ - (fun (a : Keyword) (b : (list (Vector.t bool (lambda + lambda) * Identifier lambda) * - list nat)) => - map (fun v => Dec (F x a) (fst (splitVector lambda lambda v))) (fst (split (fst b))) = (snd (split (fst b))))) in H4. - - eapply evalDist_ret_eq. - f_equal. - - eapply (@compMap_support _ _ - (fun (a : Blist) (b : Tag * (list (Vector.t bool (lambda + lambda) * Identifier lambda))) => snd b = arrayLookupList _ (combine (toW d) l0) a)) - in H7. - - eapply (@map_ext_pred _ _ _ (fun a b => a = b /\ (map (fun v => Dec (F x (fst (fst b))) (fst (splitVector lambda lambda v))) (fst (split (snd (snd b))))) = (snd (split (snd (snd b)))))). - - eapply list_pred_combine_same'. - simpl. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_same_l. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_fst_split_irr_if. - eapply H7. - intuition. - eapply H9. - intuition. - simpl in *. - subst. - pairInv. - trivial. - simpl in *. - subst. - pairInv. - - assert (l0 = fst (split x3)). - rewrite <- Heqz. - trivial. - subst. - - Theorem list_pred_arrayLookupList : - forall (A B : Set){eqda : EqDec A}(P : A -> list B -> Prop) lsa lsb, - list_pred P lsa lsb -> - forall a, - (P a nil) -> - P a (arrayLookupList _ (combine lsa lsb) a). - - induction 1; intuition; simpl in *. - unfold arrayLookupList. - simpl. - case_eq (eqb a a1); intuition. - rewrite eqb_leibniz in H2; subst. - trivial. - specialize (IHlist_pred a). - intuition. - - Qed. - - Show. - specialize (@list_pred_arrayLookupList _ _ _ - (fun a b => - (foreach (v in fst (split b)) - Dec (F x a) (fst (splitVector lambda lambda v))) = - snd (split b)) (toW d) (fst (split x3))); intuition. - eapply H6. - split_eq. - eapply H4. - simpl. - trivial. - trivial. - - (* proof of equality of GenTrans given list predicate*) - - intuition. - unfold GenTrans_3_1, GenTrans_4. - destruct b0. - destruct p1. - simpl in *. - pairInv. - comp_simp. - f_equal. - simpl in *. - eapply (@map_ext_pred _ _ _ (fun a b => a = fst b /\ (snd b = Dec (F x a1) (fst b)))). - Theorem list_pred_getSomes : - forall (A B : Set) (P : A -> B -> Prop) lsa lsb, - list_pred (fun a b => - match a with - | Some x => - match b with - | Some y => P x y - | None => False - end - | None => - match b with - | Some y => False - | None => True - end - end) lsa lsb -> - list_pred P - (getSomes lsa) (getSomes lsb). - - induction lsa; inversion 1; intuition; subst; simpl in *. - econstructor. - destruct a. - destruct a2. - inversion H. - subst; clear H. - econstructor; intuition. - - inversion H; clear H; subst. - intuition. - - destruct a2. - inversion H; clear H; subst. - intuition. - eapply IHlsa. - intuition. - - Qed. - eapply list_pred_getSomes. - eapply list_pred_map_l'. - - - Theorem list_pred_unary : - forall (A : Set)(P : A -> Prop)(ls : list A), - (forall a, In a ls -> P a) -> - list_pred (fun a b => P a /\ a = b) ls ls. - - induction ls; intuition; simpl in *. - econstructor. - econstructor; intuition. - - Qed. - - Theorem ServerSearch_4_spec : - forall a b ls z, - In z (ServerSearch_4 a b ls) -> - match z with - | None => True - | Some p => - exists e, - In e ls /\ snd e = snd p - /\ (fst (splitVector lambda lambda (fst e))) = fst p - end. - - intuition. - unfold ServerSearch_4 in *. - eapply in_map_iff in H. - destruct H. - intuition. - unfold ServerSearchLoop_4 in *. - destruct x. - destruct p0. - remember (splitVector lambda lambda b0) as y. - destruct y. - destruct (in_dec (EqDec_dec nat_EqDec) (n ^ bvectorToNat t0) a). - rewrite <- H0. - exists (b0, b1). - intuition. - eapply in_combine_r. - eauto. - simpl. - rewrite <- Heqy. - trivial. - - rewrite <- H0. - trivial. - - Qed. - - Show. - eapply list_pred_impl. - eapply list_pred_unary. - eapply ServerSearch_4_spec. - intuition. - destruct a2. - intuition. - subst. - intuition. - destruct H11. - intuition. - - Theorem map_split_eq : - forall (A B : Type)(f : A -> B)(ls : list (A * B)), - map f (fst (split ls)) = snd (split ls) -> - (forall p, In p ls -> f (fst p) = snd p). - - induction ls; intuition; simpl in *; - intuition. - - subst. - remember (split ls) as z. - destruct z. - destruct p0. - simpl in *. - inversion H; clear H; subst. - intuition. - - remember (split ls) as z. - destruct z. - destruct a. - simpl in *. - inversion H; clear H; subst. - intuition. - - Qed. - - symmetry. - destruct p0; simpl in *. - subst. - - specialize (@map_split_eq _ _ (fun (a : Bvector (lambda + lambda)) => Dec (F x a1) (fst (splitVector lambda lambda a)))); intuition. - eapply H6. - eauto. - trivial. - intuition; subst. - intuition. - intuition. - subst. - intuition. - intuition. - repeat simp_in_support. - trivial. - intuition. - - - (* correct proof of second goal begins*) - intuition. - unfold Initialize_wLoop_4 in *. - repeat simp_in_support. - - eapply (@compMap_support _ _ (fun a b => Dec (F x a1) (fst (splitVector lambda lambda (fst b))) = snd b)) in H11. - - simpl. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - split_eq. - - Theorem list_pred_impl_unary : - forall (A B: Set)(P1 : A -> B -> Prop)(P2 : A -> A -> Prop)(lsa : list A)(lsb : list B), - list_pred P1 lsa lsb -> - (forall a b, P1 a b -> P2 a a) -> - list_pred P2 lsa lsa. - - induction 1; intuition; simpl in *. - econstructor. - econstructor; intuition; eauto. - - Qed. - - eapply list_pred_impl_unary . - eapply list_pred_symm. - eapply H11. - intuition. - intuition. - unfold Initialize_indLoop_4 in *. - repeat simp_in_support. - simpl. - rewrite splitVector_append. - simpl. - eapply Enc_correct; intuition. - (* correct proof of second goal ends*) - - Qed. - - - Theorem G3_1_equiv : - Pr[G3] == Pr[G_gen Initialize_3_1]. - - unfold G3, G_gen, Initialize_3, Initialize_3_1. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_indLoop_4_eq_spec : - forall d x x1 x2 b0, - comp_spec - (fun a b => fst b = a) - (Initialize_indLoop_2 x1 x2 x b0 d) - (Initialize_indLoop_4 x1 x2 x b0 d). - - intuition. - unfold Initialize_indLoop_2. - unfold Initialize_indLoop_4. - comp_skip; try - eapply (oneVector (lambda + lambda)). - eapply comp_spec_ret; intuition. - Qed. - - Theorem Init_wLoop_4_eq_spec : - forall d x x1 x2 b0, - comp_spec - (fun a b => (fst (split (fst b)) = fst a) /\ (snd a = snd b)) - (Initialize_wLoop_2 d x x1 x2 b0) - (Initialize_wLoop_4 d x x1 x2 b0). - - intuition. - unfold Initialize_wLoop_2, Initialize_wLoop_4. - comp_simp. - comp_skip. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - eapply Init_indLoop_4_eq_spec. - eapply comp_spec_ret; intuition. - simpl. - split_eq. - eapply list_pred_symm. - eauto. - - Qed. - - eapply Init_wLoop_4_eq_spec. - remember (split a1) as z. - destruct z. - remember (split b0) as z. - destruct z. - - assert (map - (fun - v : list - (Vector.t bool (lambda + lambda) * Identifier lambda) => - fst (split v)) l2 = l0). - split_eq. - eapply list_pred_map_l'. - assert (l0 = fst (split a1)). - rewrite <- Heqz. - trivial. - subst. - assert (l2 = (fst (split b0))). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_symm. - eapply list_pred_impl. - eapply H6. - intuition. - - rewrite H7. - eapply comp_spec_seq_eq; eauto with inhabited. - eapply comp_spec_eq_refl. - intros. - comp_simp. - - comp_skip. - eapply (@compMap_spec _ _ _ _ _ _ eq - (fun a b => fst a = fst b /\ - snd a = fst (split (snd b))) - ). - eapply list_pred_eq. - intuition; subst. - comp_skip. - eapply comp_base_exists; intuition. - eapply comp_base_exists; intuition. - - eapply comp_spec_ret; intuition. - simpl. - - unfold arrayLookupList. - - Theorem arrayLookup_Some_list_pred : - forall (A B C : Set){eqda : EqDec A}(P : B -> C -> Prop)(lsa : list A) lsb lsc a b, - list_pred P lsb lsc -> - arrayLookup _ (combine lsa lsb) a = Some b-> - (exists c, - arrayLookup _ (combine lsa lsc) a = Some c /\ - P b c). - - induction lsa; intuition; simpl in *. - discriminate. - destruct lsb. - inversion H; clear H; subst. - simpl in *. - discriminate. - - inversion H; clear H; subst. - simpl in *. - case_eq (eqb a0 a); intuition. - rewrite eqb_leibniz in H. - subst. - rewrite eqb_refl in *. - inversion H0; clear H0; subst. - econstructor; eauto. - - rewrite H in H0. - eapply IHlsa; eauto. - - Qed. - - Theorem arrayLookup_None_combine : - forall (A B C : Set){eqd : EqDec A}(lsa : list A)(lsb : list B)(lsc : list C) a, - arrayLookup _ (combine lsa lsb) a = None -> - length lsb = length lsc -> - arrayLookup _ (combine lsa lsc) a = None. - - induction lsa; intuition; simpl in *. - destruct lsc; simpl in *. - trivial. - destruct lsb; simpl in *. - omega. - destruct ( eqb a0 a). - discriminate. - eauto. - - Qed. - - case_eq (@arrayLookup Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (@list_EqDec bool bool_EqDec) - (@combine Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (toW d) l2) b1); intuition. - eapply (@arrayLookup_Some_list_pred _ _ _ _ (fun a b => fst (split a) = b)) in H12. - destruct H12. - intuition. - rewrite H13. - intuition. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst. - trivial. - - eapply arrayLookup_None_combine in H12. - rewrite H12. - trivial. - rewrite map_length. - trivial. - - assert (l1 = l3). - assert (l3 = snd (split b0)). - rewrite <- Heqz0. - trivial. - assert (l1 = snd (split a1)). - rewrite <- Heqz. - trivial. - subst. - split_eq. - eapply list_pred_impl. - eauto. - intuition. - subst. - - eapply comp_spec_ret; intuition. - f_equal. - - - eapply map_ext_pred. - eapply list_pred_combine_simple. - eapply list_pred_eq. - eauto. - eapply compMap_length in H9. - rewrite H9. - symmetry. - eapply split_length_l. - - intuition. - simpl in *. - subst. - unfold GenTrans_3_1. - destruct b2. - simpl. - destruct p0. - destruct q. - comp_simp. - simpl. - - assert ( - (ServerSearch_2 (XSetSetup_2 x0 x1 d l3) - (map - (fun c : nat => - groupExp g (mult (F_P x0 k0) (F_P x2 (app k c)))) - (allNatsLt (length (fst (split l0))))) - (fst (split l0))) = -(map - (fun x3 : option (prod (Vector.t bool lambda) (Bvector lambda)) => - match x3 with - | Some v => Some (fst v) - | None => None - end) - (ServerSearch_4 (XSetSetup_2 x0 x1 d l3) - (map - (fun c : nat => - groupExp g (mult (F_P x0 k0) (F_P x2 (app k c)))) - (allNatsLt (length l0))) l0)) - - ). - - eapply list_pred_eq_impl_eq. - eapply list_pred_map_r'. - - unfold ServerSearch_2, ServerSearch_4. - eapply list_pred_map_both'. - eapply list_pred_impl. - eapply (@list_pred_combine_simple _ _ _ _ eq (fun a b => fst b = a)). - assert (length (fst (split l0)) = length l0). - eapply split_length_l. - rewrite H8. - eapply list_pred_eq. - split_eq. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst; trivial. - rewrite map_length. - rewrite allNatsLt_length. - trivial. - - intuition. - simpl in *; subst. - unfold ServerSearchLoop_4. - destruct b3. - destruct p0. - simpl. - remember (splitVector lambda lambda b2) as z. - destruct z. - destruct (in_dec (EqDec_dec nat_EqDec) (n ^ bvectorToNat t3) - (XSetSetup_2 x0 x1 d l3)); trivial. - - f_equal. - eapply map_ext_pred. - eapply (@list_pred_getSomes _ _ eq). - eapply list_pred_impl. - eapply list_pred_eq_gen. - eapply H8. - intuition; subst. - destruct b2; intuition. - intuition. - subst. - trivial. - - f_equal. - f_equal. - rewrite split_length_l. - trivial. - eapply H8. - Qed. - - (* Step 5: prepare to the replace the first PRF with a random function. *) - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - Notation "'query' v" := (OC_Query _ v) (at level 79) : comp_scope. - - Definition Initialize_wLoop_5 db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$ query w; - indLoopRes <--$$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition Initialize_5 (db : DB)(q : list Query) := - - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - k_Z <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_5 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G5_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_5 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - Definition G5 := - k_S <-$ {0,1}^lambda; - [b, _] <-$2 G5_A _ _ (f_oracle F _ k_S) tt; - ret b. - - - (* Step 4.1: move the key sampling to the beginning of the game. *) - Definition Initialize_4_1 k_S (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_4 db k_S k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G4_1 := - k_S <-$ {0,1}^lambda; - z <-$ A1; - [db, q, s_A]<-3 z; - z0 <-$ Initialize_4_1 k_S db q; - [edb, ls]<-2 z0; - A2 s_A edb (snd (split ls)). - - Theorem G4_1_equiv : - Pr[G4] == Pr[G4_1]. - - unfold G4_1, G4, G_gen. - comp_swap_r. - comp_skip. - comp_simp. - unfold Initialize_4, Initialize_4_1. - repeat (inline_first; comp_skip; try reflexivity). - Qed. - - Theorem compMap_oc_spec : - forall (C D : Set)(P2 : C -> D -> Prop)(A B : Set)(P1 : A -> B -> Prop)(eqdc : EqDec C)(eqdd : EqDec D)(E F S: Set)(eqds : EqDec S)(ls1 : list A)(ls2 : list B)(c1 : A -> Comp C)(c2 : B -> OracleComp E F D)o (s : S), - list_pred P1 ls1 ls2 -> - (forall a b z, P1 a b -> comp_spec (fun x y => P2 x (fst y)) (c1 a) (c2 b _ _ o z)) -> - comp_spec (fun a b => list_pred P2 a (fst b)) - (compMap _ c1 ls1) - ((oc_compMap _ c2 ls2) _ _ o s). - - induction ls1; inversion 1; subst; intuition; simpl in *. - comp_simp. - eapply comp_spec_ret; simpl; econstructor. - - simpl. - comp_skip. - comp_simp. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - econstructor; eauto. - - Qed. - - Theorem G5_equiv : - Pr[G4_1] == Pr[G5]. - - unfold G4_1, G5, G5_A, Initialize_4, Initialize_5. - comp_skip. - - eapply comp_spec_eq_impl_eq. - simpl. - inline_first. - comp_skip. - comp_simp. - unfold Initialize_4_1. - do 3 (simpl; inline_first; eapply comp_spec_seq_eq; eauto with inhabited; try reflexivity; intuition; comp_simp). - inline_first. - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - - unfold Initialize_wLoop_4, Initialize_wLoop_5. - simpl; inline_first. - - do 2 (simpl; inline_first; eapply comp_spec_seq_eq; eauto with inhabited; try reflexivity; intuition; comp_simp). - eapply comp_spec_ret; intuition. - - destruct b1. - simpl in *. - assert (a3 = l). - eapply list_pred_eq_impl_eq. - trivial. - subst. - remember (split l) as z. - destruct z. - do 2 (simpl; inline_first; comp_skip; comp_simp). - simpl. - comp_simp. - inline_first. - comp_simp. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - simpl; inline_first. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - (* Step 6: replace the PRF with a random function. *) - Definition G6 := - [b, _] <-$2 G5_A _ _ (@randomFunc Blist (Bvector lambda) (Rnd lambda) _) nil; - ret b. - - Theorem G6_close : - | Pr[G5] - Pr[G6] | == PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ G5_A. - - reflexivity. - - Qed. - - (* Step 7: replace random function outputs with random values. *) - Definition Initialize_wLoop_7 db k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <-$ {0, 1}^lambda; - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_7 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_7 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G7 := G_gen Initialize_7. - - (* Step 6.1: We'll use a standard argument about mapping a randomf function over a list with -no duplicates. Start by changing the structure to match that argument. *) - Definition Initialize_wLoop_6_1 db k_E k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_6_1 (db : DB)(q : list Query) := - - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - [wLoopRes, _] <-$2 - (oc_compMap _ (fun w => k_E <--$ query w; $ Initialize_wLoop_6_1 db k_E k_I k_Z w) (toW db)) - _ _ (@randomFunc Blist (Bvector lambda) (Rnd lambda) _) nil; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G6_1 := - [db, q, s_A] <-$3 A1; - z0 <-$ Initialize_6_1 db q; - [edb, ls]<-2 z0; - A2 s_A edb (snd (split ls)). - - Theorem G6_1_equiv : - Pr[G6] == Pr[G6_1]. - - unfold G6, G6_1, Initialize_6_1, G5_A. - eapply comp_spec_eq_impl_eq. - simpl; inline_first; comp_skip; comp_simp. - unfold Initialize_5. - do 3 (simpl; inline_first; comp_skip; comp_simp). - - simpl; inline_first. - comp_skip. - - Theorem oc_compMap_eq : - forall (A B C D : Set){eqd : EqDec D}(f1 f2 : A -> OracleComp B C D)(ls : list A) (S : Set){eqds : EqDec S} o (s : S), - (forall s a, comp_spec - eq - ((f1 a) _ _ o s) - ((f2 a) _ _ o s)) -> - comp_spec - eq - ((oc_compMap _ f1 ls) _ _ o s) - ((oc_compMap _ f2 ls) _ _ o s). - - induction ls; intuition; simpl in *. - comp_simp. - eapply comp_spec_eq_refl. - comp_skip. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - - Qed. - - eapply oc_compMap_eq; intuition. - unfold Initialize_wLoop_5, Initialize_wLoop_6_1. - simpl; inline_first. - - assert (comp_spec eq - - (Bind (RndPerm (length (lookupInds a a0))) - (fun a1 : list nat => - Bind - (Ret - (EqDec_dec - (pair_EqDec - (comp_EqDec (RndPerm (length (lookupInds a a0)))) - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (Bvector_EqDec lambda))))) - (pair a1 s)) - (fun z0 : prod (list nat) (list (prod Blist (Bvector lambda))) => - let 'pair z s' := z0 in - Bind (randomFunc (Rnd lambda) (list_EqDec bool_EqDec) s' a0) - (fun - z1 : prod (Bvector lambda) - (list (prod Blist (Bvector lambda))) => - let 'pair z2 s'0 := z1 in - Bind - (Bind - (compMap - (pair_EqDec - (Bvector_EqDec (plus lambda lambda)) - (Bvector_EqDec lambda)) - (Initialize_indLoop_4 b2 b3 z2 a0) - (combine - (permute - (lookupInds a a0) z) - (allNatsLt (length (lookupInds a a0))))) - (fun - x : list - (prod - (Vector.t bool (plus lambda lambda)) - (Identifier lambda)) => - Ret - (EqDec_dec - (pair_EqDec - (comp_EqDec - (compMap - (pair_EqDec - (Bvector_EqDec - (plus lambda lambda)) - (Bvector_EqDec lambda)) - (Initialize_indLoop_4 b2 b3 z2 - a0) - (combine - (permute - - (lookupInds a a0) z) - (allNatsLt - (length (lookupInds a a0)))))) - (list_EqDec - (pair_EqDec - (list_EqDec bool_EqDec) - (Bvector_EqDec lambda))))) - (pair x s'0))) - (fun - z3 : prod - (list - (prod - (Vector.t bool - (plus lambda lambda)) - (Identifier lambda))) - (list (prod Blist (Bvector lambda))) => - let 'pair z4 s'1 := z3 in - Bind - (Ret - (EqDec_dec - (pair_EqDec - (list_EqDec - (pair_EqDec - (Bvector_EqDec - (plus lambda lambda)) - (Bvector_EqDec lambda))) - (list_EqDec nat_EqDec))) - (pair z4 z)) - (fun - x : prod - (list - (prod - (Vector.t bool - (plus lambda lambda)) - (Identifier lambda))) - (list nat) => - Ret - (EqDec_dec - (pair_EqDec - (comp_EqDec - (Ret - (EqDec_dec - (pair_EqDec - (list_EqDec - (pair_EqDec - (Bvector_EqDec - (plus lambda lambda)) - (Bvector_EqDec lambda))) - (list_EqDec nat_EqDec))) - (pair z4 z))) - (list_EqDec - (pair_EqDec - (list_EqDec bool_EqDec) - (Bvector_EqDec lambda))))) - (pair x s'1))))))) - - (a1 <-$ RndPerm (length (lookupInds a a0)); - z1 <-$ randomFunc ({ 0 , 1 }^lambda) (list_EqDec bool_EqDec) s a0; - [z2, s'0]<-2 z1; - z3 <-$ - (x <-$ - (foreach (x - in combine (permute (lookupInds a a0) a1) - (allNatsLt (length (lookupInds a a0)))) - Initialize_indLoop_4 b2 b3 z2 a0 x); ret (x, s'0)); - [z4, s'1]<-2 z3; x <-$ ret (z4, a1); ret (x, s'1))). - - - eapply eq_impl_comp_spec_eq. - intuition. - comp_skip. - comp_skip. - comp_simp. - inline_first. - comp_skip. - reflexivity. - comp_simp. - eapply evalDist_ret_eq. - trivial. - rewrite H1. - clear H1. - - eapply eq_impl_comp_spec_eq. - intros. - comp_swap_l. - comp_skip. - reflexivity. - comp_simp. - eapply comp_spec_eq_impl_eq. - simpl; inline_first. - comp_skip. - simpl; inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - - subst. - remember (split a1) as z. - destruct z. - - do 2 (simpl; inline_first; comp_skip; comp_simp). - inline_first; comp_simp. - simpl. - inline_first. - eapply comp_spec_eq_trans_l. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - - Qed. - - Definition Initialize_6_2 (db : DB)(q : list Query) := - - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - [wLoopRes, _] <-$2 - (oc_compMap _ (fun w => k_E <--$ query w; $ Initialize_wLoop_6_1 db k_E k_I k_Z w) (toW db)) - _ _ (fun (s: unit)(a : Blist) => b <-$ {0, 1}^lambda; ret (b, s)) tt; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G6_2 := - [db, q, s_A] <-$3 A1; - z0 <-$ Initialize_6_2 db q; - [edb, ls]<-2 z0; - A2 s_A edb (snd (split ls)). - - Theorem G6_2_equiv : - Pr[G6_1] == Pr[G6_2]. - - unfold G6_1, G6_2. - comp_skip. - comp_simp. - unfold Initialize_6_1, Initialize_6_2. - do 3 (inline_first; comp_skip; comp_simp). - inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip. - - Theorem compMap_randomFunc_NoDup : - forall (A B C: Set){eqda : EqDec A}{eqdb : EqDec B}{eqdc : EqDec C}(ls : list A)(f : A -> B -> Comp C)(rndB : Comp B)(lsf : list (A * B)), - NoDup ls -> - (forall a, In a ls -> arrayLookup _ lsf a = None) -> - comp_spec (fun a b => fst a = fst b) - ((oc_compMap _ (fun x => y <--$ query x; $ f x y) ls) _ _ (@randomFunc A B rndB _) lsf) - ((oc_compMap _ (fun x => y <--$ query x; $ f x y) ls) _ _ (fun s a => b <-$ rndB; ret (b, s)) tt). - - induction ls; intuition; simpl in *. - comp_simp. - eapply comp_spec_ret; intuition. - inversion H; clear H; subst. - - simpl; inline_first. - unfold randomFunc. - rewrite H0. - inline_first. - comp_skip. - comp_simp. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply IHls; intuition. - simpl. - case_eq (eqb a0 a); intuition. - rewrite eqb_leibniz in H7. - subst. - intuition. - comp_simp. - simpl in *; subst. - eapply comp_spec_ret; intuition. - intuition. - - Qed. - - eapply compMap_randomFunc_NoDup. - eapply removeDups_NoDup. - intuition. - - destruct b0. - simpl in *; subst. - remember (split l0) as z. - destruct z. - inline_first. - reflexivity. - - Qed. - - Theorem G7_equiv_h : - Pr[G6_2] == Pr[G7]. - - unfold G6_2, G7, G_gen, Initialize_6_2, Initialize_7. - - do 4 (comp_skip; comp_simp; inline_first). - eapply comp_spec_eq_impl_eq. - - comp_skip. - eapply comp_spec_symm. - eapply (@compMap_oc_spec _ _ (fun a b => a = b)). - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_7, Initialize_wLoop_6_1. - simpl; comp_simp. - inline_first. - comp_swap_l. - do 2 (comp_skip; comp_simp; inline_first). - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - - simpl in *. - remember (split a1) as z. - destruct z. - remember (split b0) as z. - destruct z. - assert (l0 = l2). - assert (l0 = fst (split a1)). - rewrite <- Heqz. - trivial. - assert (l2 = fst (split b0)). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_symm. - eapply list_pred_impl. - eapply H6. - intuition. - subst; trivial. - subst. - - inline_first; comp_skip. - inline_first; comp_skip. - comp_simp. - assert (l1 = l3). - assert (l1 = snd (split a1)). - rewrite <- Heqz. - trivial. - assert (l3 = snd (split b0)). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_symm. - eapply list_pred_impl. - eapply H6. - intuition. - subst; trivial. - subst. - reflexivity. - Qed. - - - (* Step 8: Apply the semantic security definition to replace the ciphertexts with encryptions of 0^lambda *) - Definition zeroVector n := Vector.const false n. - - Definition Initialize_indLoop_8 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_8 db k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <-$ {0, 1}^lambda; - indLoopRes <-$ compMap _ (Initialize_indLoop_8 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_8 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_8 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G8 := G_gen Initialize_8. - - Require Import Encryption. - - (* Step 7.1: We will use an iterated semantic security definition, and we need to get the game in the correct form. Start by moving some things around. *) - - Definition Initialize_indLoop_7_1 (enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda)) k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ enc k_E ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_7_1 (enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda)) db k_I k_Z w := - k_E <-$ {0, 1}^lambda; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - indLoopRes <-$ compMap _ (Initialize_indLoop_7_1 enc k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_7_1 (enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda)) (db : DB)(q : list Query) := - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_7_1 enc db k_I k_Z) (toW db); - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - - Definition G7_1 (enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda)) := - [db, q, s_A] <-$3 A1; - z0 <-$ Initialize_7_1 enc db q; - [edb, ls]<-2 z0; - A2 s_A edb (snd (split ls)). - - - Theorem G7_1_equiv : - Pr[G7] == Pr[G7_1 Enc]. - - unfold G7, G7_1, Initialize_7, Initialize_7_1. - comp_skip. - comp_simp. - do 3 (inline_first; comp_at comp_inline leftc 1%nat; comp_swap_l; comp_skip). - - eapply compMap_eq. - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_7, Initialize_wLoop_7_1. - comp_simp. - comp_swap_r. - reflexivity. - - inline_first. - comp_skip. - - reflexivity. - - Qed. - - (* Step 7.2: Put the game in the form of the encryption definition. *) - Definition Initialize_indLoop_7_2 k_I k_Z w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$ query ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - $ ret (@Vector.append bool lambda lambda e (natToBvector y), ind). - - - Definition Initialize_wLoop_7_2 db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_7_2 k_I k_Z w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition EncI_A1 := - [db, q, s_A] <-$3 A1; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - ret (toW db, (k_I, k_Z, db, q, s_A)). - - Definition EncI_A2 (s : Bvector lambda * Bvector lambda * DB * list Query * A_State) w := - let '(k_I, k_Z, db, q, s_A) := s in - Initialize_wLoop_7_2 db k_I k_Z w. - - Definition EncI_A3 (s : Bvector lambda * Bvector lambda * DB * list Query * A_State) wLoopRes := - let '(k_I, k_Z, db, q, s_A) := s in - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - Definition G7_2 (enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda)) := - [ws, s] <-$2 EncI_A1; - wLoopRes <-$ compMap _ (fun x => - key <-$ {0,1}^lambda; - [b, _] <-$2 (EncI_A2 s x) _ _ (EncryptOracle enc _ key) tt; - ret b) ws; - EncI_A3 s wLoopRes. - - Theorem G7_2_equiv : forall enc, - Pr[G7_1 enc] == Pr[G7_2 enc]. - - intuition. - unfold G7_2, G_gen, EncI_A1. - inline_first. - comp_skip. - comp_simp. - unfold Initialize_7_1. - inline_first. - comp_skip. - inline_first. - comp_skip. - comp_simp. - inline_first. - comp_skip. - eapply compMap_eq. - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_7_1. - inline_first. - comp_skip. - unfold EncI_A2. - unfold Initialize_wLoop_7_2. - comp_simp. - eapply comp_spec_eq_impl_eq. - simpl. - inline_first. - comp_skip. - comp_simp. - inline_first. - comp_skip. - eapply (@compMap_oc_spec _ _ (fun a b => a = b)). - eapply list_pred_eq. - intuition. - pairInv. - unfold Initialize_indLoop_4, Initialize_indLoop_7_2. - simpl. - unfold EncryptOracle. - inline_first. - comp_skip; try - eapply (oneVector (lambda + lambda)). - comp_simp. - eapply comp_spec_ret; intuition. - - comp_simp. - inline_first. - comp_simp. - eapply comp_spec_ret; intuition. - f_equal. - simpl in *. - eapply list_pred_eq_impl_eq; intuition. - - unfold EncI_A3. - inline_first. - comp_skip. - remember (split x1) as z. - destruct z. - - remember ( - @split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1 - ) as z. - destruct z. - inline_first. - assert (l0 = l2). - assert (l0 = (fst (split x1))). - rewrite <- Heqz. - trivial. - assert (l2 = (fst (@split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1))). - rewrite <- Heqz0. - trivial. - subst. - trivial. - - subst. - comp_skip. - reflexivity. - comp_simp. - inline_first. - comp_skip. - reflexivity. - - assert (l1 = l3). - assert (l1 = (snd (split x1))). - rewrite <- Heqz. - trivial. - assert (l3 = (snd ((@split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1)))). - rewrite <- Heqz0. - trivial. - subst. - trivial. - subst. - reflexivity. - - Qed. - - - (* Step 7.3 : Apply the encryption security definition. *) - Definition G7_3 := - [ws, s] <-$2 EncI_A1; - wLoopRes <-$ compMap _ (fun x => - key <-$ {0,1}^lambda; - [b, _] <-$2 (EncI_A2 s x) _ _ (EncryptNothingOracle Enc _ (zeroVector lambda) key) tt; - ret b) ws; - EncI_A3 s wLoopRes. - - Theorem G7_3_close : - | Pr[G7_2 Enc] - Pr[G7_3] | <= IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc EncI_A1 EncI_A2 EncI_A3 _ _ (zeroVector lambda). - - eapply leRat_refl. - - Qed. - - (* Now put everything back. Start by showing that the last game is the same as G7_2 with a different encrypt function. *) - Theorem G7_3_back_equiv : - Pr[G7_3] == Pr[G7_2 (fun k p => Enc k (zeroVector lambda))]. - - - reflexivity. - - Qed. - - - (* use G7_2_equiv to do the next step *) - - Theorem G8_equiv_h : - Pr[G7_1 (fun k p => Enc k (zeroVector lambda))] == Pr[G8]. - - unfold G8, G_gen, G7_1, Initialize_7_1, Initialize_8. - comp_skip; comp_simp. - do 3 (inline_first; comp_at comp_inline rightc 1%nat; comp_swap_r; comp_skip). - - eapply compMap_eq. - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_7_1, Initialize_wLoop_8. - comp_simp. - comp_swap_l. - reflexivity. - - inline_first; comp_skip. - reflexivity. - - Qed. - - - (* Step 9: Next we will replace the "Z" PRF with a random function. Start by putting the game in the oracle form. *) - Definition Initialize_indLoop_9 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ query (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_9 db k_I w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_9 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition GenTrans_9 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret g^(e * z)) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_9 (db : DB)(q : list Query) := - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_9 db k_I) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_9 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G9_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_9 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - Definition G9 := - k <-$ {0, 1}^lambda; - [b, _] <-$2 G9_A _ _ (f_oracle F_P _ k) tt; - ret b. - - - (* Step 8.1: Start by sampling the key at the beginning. *) - Definition G8_1 := - - k_Z <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_8 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - - Theorem G8_1_equiv : - Pr[G8] == Pr[G8_1]. - - unfold G8, G8_1, G_gen, Initialize_8. - comp_swap_r. - comp_skip. - comp_simp. - inline_first. - comp_swap_r. - comp_skip. - inline_first. - comp_swap_r. - comp_skip. - - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - - Qed. - - - (* Step 8.2: put some functions in the form of computations. *) - - Definition GenTrans_8_2 (edb : EDB) k_X k_Z (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <-$ compMap _ (fun (c : nat) => ret g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - ret (inds, (stag, (combine xtokens res))). - - Definition G8_2 := - - k_Z <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_8 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <-$ compMap _ (GenTrans_8_2 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - - Theorem G8_2_equiv : - Pr[G8_1] == Pr[G8_2]. - - unfold G8_1, G8_2, G_gen, Initialize_8. - do 6 (comp_skip; comp_simp). - comp_skip. - comp_irr_r. - eapply compMap_wf; intuition. - unfold GenTrans_8_2. - comp_simp. - wftac. - eapply compMap_wf; intuition. - wftac. - eapply (@compMap_support _ _ (fun a b => b = GenTrans_4 (t, XSetSetup_2 x0 x1 d l1) x0 x a)) in H8. - assert (map ( GenTrans_4 (t, XSetSetup_2 x0 x1 d l1) x0 x) (combine l x3) = x4). - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eauto. - intuition. - subst. - reflexivity. - - intuition. - unfold GenTrans_8_2, GenTrans_4 in *. - comp_simp. - repeat simp_in_support. - assert (x5 = map (fun (c : nat) => g ^ (F_P x0 k0 * F_P x (k ++ c))) (allNatsLt (length b0))). - eapply list_pred_eq_impl_eq. - - eapply (@compMap_support _ _ (fun (a : nat) b => b = g ^ (F_P x0 k0 * F_P x (k ++ a)))) in H11. - eapply list_pred_symm. - eapply list_pred_map_l'. - eapply H11. - - intuition. - repeat simp_in_support. - trivial. - - subst. - reflexivity. - - Qed. - - - Theorem G9_equiv : - Pr[G8_2] == Pr[G9]. - - unfold G8_2, G9, G9_A, Initialize_9. - comp_skip. - - eapply comp_spec_eq_impl_eq. - do 3 (simpl; inline_first; comp_skip; comp_simp). - - simpl; inline_first; comp_skip. - - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_8, Initialize_wLoop_9. - comp_simp. - repeat (simpl; inline_first; comp_skip; comp_simp). - - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold Initialize_indLoop_8, Initialize_indLoop_9. - - repeat (simpl; inline_first; comp_skip; comp_simp); - try - eapply (oneVector (lambda + lambda)). - eapply comp_spec_ret; intuition. - eapply comp_spec_ret; intuition. - simpl in *. - f_equal. - eapply list_pred_eq_impl_eq; intuition. - - destruct b3. - simpl in *. - assert (a0 = l). - eapply list_pred_eq_impl_eq; intuition. - subst. - remember (split l) as z. - destruct z. - - repeat (simpl; inline_first; comp_skip; comp_simp). - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold GenTrans_8_2, GenTrans_9. - - comp_simp. - simpl. - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - simpl. - unfold f_oracle. - comp_simp. - eapply comp_spec_ret; intuition. - comp_simp. - simpl in *. - eapply comp_spec_ret; intuition. - simpl. - assert (a0 = l2). - eapply list_pred_eq_impl_eq. - intuition. - subst. - reflexivity. - - inline_first; - comp_simp. - simpl. - inline_first. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - simpl in *. - assert (a0 = l2). - eapply list_pred_eq_impl_eq. - intuition. - subst. - eapply comp_spec_eq_refl. - subst. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - (* Step 10: Replace the PRF with a random function. *) - Definition G10 := - [b, _] <-$2 G9_A _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - ret b. - - Theorem G10_close : - | Pr[G9] - Pr[G10] | <= PRF_Advantage (Rnd lambda) (RndNat (2^lambda)) F_P _ _ G9_A. - - reflexivity. - - Qed. - - (* DON'T USE PERMUTATIONS!!! Use a blinding construction that returns "None" when it is unblinded with an incompatible value. *) - - (* - (* Step 10.1: change the codomain of the random function to an "option" type, but keep all values the same. *) - Definition randomFuncOpt s d := - [r, s'] <-$2 (@randomFunc Blist _ (RndNat (2^lambda)) _ s d); - ret (Some r, s'). - - Definition Initialize_indLoop_10_1 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - opt_z <--$ query (w ++ c); - $ ret - match opt_z with - | Some z => - y <- xind * (inverse z); - Some (Vector.append e (natToBvector y), ind) - | None => - None - end. - - Fixpoint hasNone(A : Type)(ls : list (option A)) := - match ls with - | nil => false - | o :: ls' => - match o with - | None => true - | Some _ => hasNone ls' - end - end. - - Definition Initialize_wLoop_10_1 db k_I e := - [w, sigma] <-2 e; - inds <- lookupInds db w; - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_10_1 k_I k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))); - $ ret - if(hasNone indLoopRes) then None else Some (getSomes indLoopRes, sigma). - - Definition computeToken_10_1 e s (c : nat) := - opt_z <--$ query (s ++ c); - $ ret - match opt_z with - | None => None - | Some z => - Some (g ^ (e * z)) - end. - - Definition GenTrans_10_1 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens_raw <--$ oc_compMap _ (computeToken_10_1 e s) (allNatsLt (length t)); - if (hasNone xtokens_raw) then $ ret None else - xtokens <- getSomes xtokens_raw; - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (Some (inds, (stag, (combine xtokens res)))). - - - Definition sampleSigmas_11 db w := - inds <- lookupInds db w; - RndPerm (length inds). - - - Definition Initialize_10_1 k_X k_I (db : DB)(q : list Query) sigmas := - xSet <- XSetSetup_2 k_X k_I db sigmas; - wLoopRes_raw <--$ oc_compMap _ (Initialize_wLoop_10_1 db k_I) (combine (toW db) sigmas); - if (hasNone wLoopRes_raw) then (OC_Ret _ _ (ret None)) else - wLoopRes <- getSomes wLoopRes_raw; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes_raw <--$ oc_compMap _ (GenTrans_10_1 edb k_X) (combine q stags_ts); - $ ret - if (hasNone searchRes_raw) then None else - searchRes <- getSomes searchRes_raw; - (Some (edb, searchRes)). - - Definition G10_1_A k_X k_I db q s_A sigmas := - z0_raw <--$ Initialize_10_1 k_X k_I db q sigmas; - match z0_raw with - | None => $ ret false - | Some z0 => - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)) - end. - - Definition G10_1 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - [b, _] <-$2 (G10_1_A k_X k_I db q s_A sigmas) _ _ (@randomFuncOpt) nil; - ret b. - - *) - - (* Step 11: Keep track of the blinding values and make sure nothing matches when we use the wrong value to unblind. *) - Definition Initialize_indLoop_11 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ query (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y), (ind, z)). - - Definition Initialize_wLoop_11 db k_I e := - [w, sigma] <-2 e; - inds <- lookupInds db w; - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_11 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition ServerSearchLoop_11 xSet (e : (nat * nat) * (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - [xtoken_z1, t_c_ind_z2] <-2 e; - [xtoken, z1] <-2 xtoken_z1; - [t_c, ind_z2] <-2 t_c_ind_z2; - [ind, z2] <-2 ind_z2; - if (eq_nat_dec z1 z2) then - ( - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some (e, ind)) else None - ) - else None. - - Definition ServerSearch_11 xSet (xtokens : list (nat * nat)) (t : list (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - map (ServerSearchLoop_11 xSet) (combine xtokens t). - - Definition GenTrans_11 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda * nat)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret (g^(e * z), z)) (allNatsLt (length t)); - res <- ServerSearch_11 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine (fst (split xtokens)) res))). - - - - Definition Initialize_11 k_X k_I (db : DB)(q : list Query) sigmas xSet := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_11 db k_I) (combine (toW db) sigmas); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_11 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - - Definition sampleSigmas_11 db w := - inds <- lookupInds db w; - RndPerm (length inds). - - Definition G11 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, _] <-$2 (Initialize_11 k_X k_I db q sigmas xSet) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - [edb, ls] <-2 z; - A2 s_A edb (snd (split ls)). - - - (* Step 10.1 : move some things around and add some bookkeeping to get closer to game 11. *) - - Definition Initialize_indLoop_10_1 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ query (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y), (ind, z)). - - Definition Initialize_wLoop_10_1 db k_I w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_11 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition ServerSearchLoop_10_1 xSet (e : (nat * nat) * (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - [xtoken_z1, t_c_ind_z2] <-2 e; - [xtoken, z1] <-2 xtoken_z1; - [t_c, ind_z2] <-2 t_c_ind_z2; - [ind, z2] <-2 ind_z2; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some (e, ind)) else None. - - Definition ServerSearch_10_1 xSet (xtokens : list (nat * nat)) (t : list (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - map (ServerSearchLoop_10_1 xSet) (combine xtokens t). - - Definition GenTrans_10_1 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda * nat)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret (g^(e * z), z)) (allNatsLt (length t)); - res <- ServerSearch_10_1 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine (fst (split xtokens)) res))). - - Definition Initialize_10_1 k_X k_I (db : DB)(q : list Query) := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_10_1 db k_I) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_10_1 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G10_1 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - [z, _] <-$2 (Initialize_10_1 k_X k_I db q) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - [edb, ls] <-2 z; - A2 s_A edb (snd (split ls)). - - Theorem G10_1_equiv : - Pr[G10] == Pr[G10_1]. - - unfold G10, G10_1, G9_A. - Local Opaque evalDist. - simpl. - inline_first. - comp_at comp_swap rightc 1%nat. - comp_swap_r. - comp_skip. - comp_simp. - simpl. - inline_first. - do 2 (comp_skip; comp_simp; inline_first). - eapply comp_spec_eq_impl_eq. - comp_skip. - - Theorem oc_compMap_spec: - forall (D E S A1 A2 : Set) P1 P2 (inv : S -> Prop) (B C : Set) (eqdd : EqDec D)(eqde : EqDec E) (f1 : A1 -> OracleComp B C D)(f2 : A2 -> OracleComp B C E) - (ls1 : list A1)(ls2 : list A2)(eqds : EqDec S) (o : S -> B -> Comp (C * S)) - (s : S) , - list_pred P1 ls1 ls2 -> - inv s -> - (forall (s0 : S) (a1 : A1)(a2 : A2), - inv s0 -> - P1 a1 a2 -> - In a1 ls1 -> - In a2 ls2 -> - comp_spec (fun a b => (P2) (fst a) (fst b) /\ snd a = snd b /\ inv (snd a)) ((f1 a1) S eqds o s0) ((f2 a2) S eqds o s0)) -> - comp_spec (fun a b => list_pred (P2) (fst a) (fst b) /\ snd a = snd b /\ inv (snd a)) ((oc_compMap _ f1 ls1) S eqds o s) - ((oc_compMap _ f2 ls2) S eqds o s). - - induction ls1; inversion 1; intuition; simpl in *; subst. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - econstructor. - - comp_skip. - comp_simp. - simpl in *. - subst. - - comp_skip. - comp_simp. - simpl in *. - eapply comp_spec_ret; intuition. - subst. - simpl. - econstructor; intuition. - - Qed. - - eapply (@oc_compMap_spec _ _ _ _ _ _ - (fun a b => - (fst (split (fst a))) = (fst (split (fst b))) /\ - (snd (split (fst a))) = (fst (split (snd (split (fst b))))) /\ - snd a = snd b - ) - (fun s => True) - ). - - eapply list_pred_eq. - intuition. - intuition; subst. - unfold Initialize_wLoop_9, Initialize_wLoop_10_1. - comp_simp. - simpl. - inline_first. - comp_skip. - comp_simp. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply (@oc_compMap_spec _ _ _ _ _ _ (fun a b => fst a = fst b /\ snd a = fst (snd b)) (fun s => True)). - eapply list_pred_eq; intuition. - intuition. - intuition. - subst. - - unfold Initialize_indLoop_9, Initialize_indLoop_11. - comp_simp; simpl. - inline_first. - comp_skip. - eapply (oneVector (lambda + lambda)). - eapply (oneVector (lambda + lambda)). - comp_simp. - comp_skip. - eapply (oneVector (lambda + lambda)). - eapply (oneVector (lambda + lambda)). - comp_simp. - eapply comp_spec_ret; intuition. - comp_simp. - simpl in *. - eapply comp_spec_ret; intuition. - simpl. - subst. - split_eq. - eapply list_pred_impl; eauto. - intuition. - subst. - simpl. - split_eq. - eapply list_pred_impl; eauto. - intuition. - - destruct b0. - remember (split a1) as z. - destruct z. - remember (split l0) as z. - destruct z. - simpl. - inline_first. - simpl in *; intuition. - subst. - assert ((foreach (v in l2)fst (split v)) = - (foreach (v in l4)fst (split v))). - - eapply (@map_ext_pred _ _ _ - (fun - (a0 : list (Vector.t bool (lambda + lambda) * Identifier lambda)) - (b : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) => - fst (split a0) = fst (split b) /\ - snd (split a0) = fst (split (snd (split b))) - ) - ). - - assert (l2 = (fst (split a1))). - rewrite <- Heqz. - trivial. - assert (l4 = fst (split l0)). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_impl; eauto. - intuition. - intuition. - - rewrite H5. - comp_skip. - simpl; inline_first. - comp_simp. - simpl. - inline_first. - eapply (@comp_spec_seq _ _); eauto with inhabited. - eapply (@compMap_spec _ _ _ _ _ _ _ (fun a b => fst a = fst b)). - eapply list_pred_eq. - - intuition; subst. - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - eapply comp_spec_ret; intuition. - intuition. - comp_simp. - inline_first. - - (* - - eapply (@compMap_support _ _ (fun a b => (snd b) = arrayLookupList (list_EqDec bool_EqDec) - (combine (toW d) l2) a)) in H10. - eapply (@compMap_support _ _ (fun a b => (snd b) = arrayLookupList (list_EqDec bool_EqDec) - (combine (toW d) l4) a)) in H11. - - comp_skip. - eapply oc_compMap_spec. - eapply list_pred_combine_l. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_eq. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_fst_split_irr_if. - eapply H11. - intuition. - eapply H13. - intuition. - simpl in *. - subst. - - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_fst_split_irr_if. - eapply H10. - intuition. - eapply H13. - - intuition. - eapply H13. - intuition. - simpl in *. - subst. -*) - - Admitted. - - (* Step 10.2 : compute the permutations and the XSet at the beginning. *) - Definition Initialize_10_2 k_X k_I (db : DB)(q : list Query) sigmas xSet := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_11 db k_I) (combine (toW db) sigmas); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_10_1 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G10_2 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, _] <-$2 (Initialize_10_2 k_X k_I db q sigmas xSet) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - [edb, ls] <-2 z; - A2 s_A edb (snd (split ls)). - - Theorem oc_compMap_fission_r : - forall (A B C D E : Set)(eqdd : EqDec D)(eqde : EqDec E)(ls : list A)(c1 : A -> OracleComp B C D)(c2 : D -> Comp E) - (S : Set)(eqds : EqDec S)(o : S -> B -> Comp (C * S))(s : S) x, - evalDist ((oc_compMap _ (fun a => d <--$ c1 a; $ (c2 d)) ls) _ _ o s) x == - evalDist ([lsD, s'] <-$2 (oc_compMap _ c1 ls) _ _ o s; v <-$ (compMap _ c2 lsD); ret (v, s')) x. - - Local Opaque evalDist. - induction ls; intuition; simpl in *. - inline_first; comp_simp. - simpl. - comp_simp. - eapply evalDist_ret_eq; intuition. - - inline_first. - comp_skip. - comp_simp. - - inline_first. - - assert ( - (evalDist - (Bind (c2 d) - (fun a1 : E => - Bind - (Ret (EqDec_dec (pair_EqDec (comp_EqDec (c2 d)) eqds)) - (pair a1 s0)) - (fun z0 : prod E S => - let 'pair z s' := z0 in - Bind - ((oc_compMap eqde - (fun a2 : A => - OC_Bind (c1 a2) (fun d0 : D => OC_Ret B C (c2 d0))) - ls) S eqds o s') - (fun z1 : prod (list E) S => - let 'pair z2 s'0 := z1 in - Bind - (Ret (EqDec_dec (list_EqDec eqde)) (cons z z2)) - (fun x : list E => - Ret - (EqDec_dec - (pair_EqDec - (comp_EqDec - (Ret (EqDec_dec (list_EqDec eqde)) - (cons z z2))) eqds)) - (pair x s'0)))))) - (pair a0 b)) - == - evalDist - (z1 <-$ - (oc_compMap eqde (fun a2 : A => d0 <--$ c1 a2; $ c2 d0) ls) S eqds o s0; - a1 <-$ c2 d; - - [z2, s'0]<-2 z1; x <-$ ret a1 :: z2; ret (x, s'0)) - (a0, b) - ). - - comp_swap_r. - comp_skip. - comp_skip. - comp_simp. - eapply evalDist_ret_eq; intuition. - rewrite H0. - clear H0. - eapply eqRat_trans. - comp_skip. - eapply eqRat_refl. - inline_first. - comp_skip. - comp_simp. - inline_first. - comp_simp. - simpl. - inline_first. - comp_at comp_inline rightc 1%nat. - comp_swap_r. - comp_skip. - Qed. - - Theorem oc_compMap_fission_l : - forall (A B C D E : Set)(eqdd : EqDec D)(eqde : EqDec E)(ls : list D)(c1 : E -> OracleComp B C D)(c2 : D -> Comp E) - (S : Set)(eqds : EqDec S)(o : S -> B -> Comp (C * S))(s : S) x, - evalDist ((oc_compMap _ (fun d => e <--$$ c2 d; (c1 e)) ls) _ _ o s) x == - evalDist ( lsE <-$ compMap _ c2 ls; (oc_compMap _ c1 lsE) _ _ o s) x. - - Local Opaque evalDist. - induction ls; intuition; simpl in *. - inline_first; comp_simp. - simpl. - comp_simp. - eapply evalDist_ret_eq; intuition. - - inline_first. - comp_skip. - inline_first. - - assert ( - (evalDist - (Bind (compMap eqde c2 ls) - (fun a1 : list E => - Bind (Ret (EqDec_dec (list_EqDec eqde)) (cons x a1)) - (fun lsE : list E => (oc_compMap eqdd c1 lsE) S eqds o s))) - (pair a0 b)) - == - evalDist - (a1 <-$ (foreach (x in ls)c2 x); - [z, s'] <-$2 (c1 x) _ _ o s; - [lsE, s''] <-$2 (oc_compMap eqdd c1 a1) S eqds o s'; - ret (z :: lsE, s'')) - (a0, b) - ). - - comp_skip. - reflexivity. - simpl. - comp_skip. - comp_simp. - comp_skip. - comp_simp. - eapply evalDist_ret_eq; intuition. - rewrite H0. - clear H0. - comp_swap_r. - comp_skip. - comp_simp. - eapply eqRat_trans. - comp_skip. - eapply eqRat_refl. - inline_first. - comp_skip. - reflexivity. - comp_skip. - comp_simp. - eapply evalDist_ret_eq. - intuition. - Qed. - - (* Step 10.3 : Check the random function for badness and expose it. *) - - Variable MaxMatchingIDs : nat. - - Print DB_Entry. - - Definition allPairs_1(A B : Type)(a : A)(lsb : list B) := - map (pair a) lsb. - - Definition allPairs_2(A B : Type)(lsa : list A)(lsb : list B) : list (A * B) := - flatten (map (fun a => allPairs_1 a lsb) lsa). - - Definition badBlinding_ezxz (xSet : list nat) (e z : nat)(p : nat * nat) := - [xind, y] <-2 p; - if (eq_nat_dec y z) then false else - if (in_dec (EqDec_dec _) (g^(e * z * xind * (inverse y))) xSet ) then true else false. - - Definition badBlinding_ez (xinds : list nat)(f : list (Blist * nat)) (xSet : list nat) (e z : nat) := - allXindsZs <- allPairs_2 xinds (snd (split f)); - fold_left (fun (b : bool) x => if b then true else badBlinding_ezxz xSet e z x) allXindsZs false. - - Definition badBlinding_c k_X (xinds : list nat)(f : list (Blist * nat)) xSet (q : Query)(c : nat) := - [s, x] <-2 q; - e <- F_P k_X x; - z_opt <- arrayLookup _ f (s ++ c); - match z_opt with - | None => false - | Some z => badBlinding_ez xinds f xSet e z - end. - - Definition badBlinding_q k_X (xinds : list nat)(f : list (Blist * nat)) xSet (q : Query) := - fold_left (fun (b : bool) z => if b then true else badBlinding_c k_X xinds f xSet q z) (allNatsLt MaxMatchingIDs) false. - - Definition badBlinding k_I k_X (db : DB)(f : list (Blist * nat)) xSet (q : list Query) := - xinds <- map (fun x => F_P k_I (Vector.to_list x)) (fst (split db)); - fold_left (fun (b : bool) x => if b then true else (badBlinding_q k_X xinds f xSet x)) q false. - - Definition goodBlinding_ezxz_Prop (xSet : list nat) (e z : nat)(xind y : nat) := - In (g ^ (e * z * xind * (inverse y))) xSet -> - y = z. - - Definition goodBlinding_ez_Prop (xinds : list nat)(f : list (Blist * nat)) (xSet : list nat) (e z : nat) := - forall xind y, - In (xind, y) (allPairs_2 xinds (snd (split f))) -> - goodBlinding_ezxz_Prop xSet e z xind y. - - Definition goodBlinding_c_Prop k_X (xinds : list nat)(f : list (Blist * nat)) xSet s x (c : nat) := - forall z, - arrayLookup _ f (s ++ c) = Some z -> - goodBlinding_ez_Prop xinds f xSet (F_P k_X x) z. - - - Definition goodBlinding_q_Prop k_X (xinds : list nat)(f : list (Blist * nat)) xSet s x := - forall z, - In z (allNatsLt MaxMatchingIDs) -> - goodBlinding_c_Prop k_X xinds f xSet s x z. - - Theorem fold_left_bad_true : - forall (A : Type)(f : A -> bool) ls, - fold_left (fun (b : bool)(a : A) => if b then true else (f a)) ls true = true. - - induction ls; intuition. - - Qed. - - Theorem badBlinding_ez_Prop_correct_h : - forall ls xSet y z0 xind y0, - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet y z0 x) - ls false = false -> - In (xind, y0) ls -> - goodBlinding_ezxz_Prop xSet y z0 xind y0. - - induction ls; intuition; simpl in *. - intuition. - intuition. - pairInv. - - destruct (eq_nat_dec y0 z0); subst. - unfold goodBlinding_ezxz_Prop. - intuition. - - destruct ( in_dec (EqDec_dec nat_EqDec) (g ^ (y * z0 * xind * inverse y0)) - xSet). - rewrite fold_left_bad_true in H. - discriminate. - - unfold goodBlinding_ezxz_Prop. - intuition. - - destruct (eq_nat_dec b z0); subst. - eapply IHls; intuition. - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (y * z0 * a0 * inverse b)) xSet). - rewrite fold_left_bad_true in H. - discriminate. - eapply IHls; intuition. - Qed. - - Theorem badBlinding_ez_Prop_correct : - forall xinds f xSet y z0, - badBlinding_ez xinds f xSet y z0 = false -> - goodBlinding_ez_Prop xinds f xSet y z0. - - unfold goodBlinding_ez_Prop. - intuition. - eapply badBlinding_ez_Prop_correct_h. - eapply H. - trivial. - - Qed. - - Theorem goodBlinding_q_Prop_correct_h : - forall ls k_X xinds f xSet s x z, - fold_left - (fun (b : bool) (z : nat) => - if b then true else badBlinding_c k_X xinds f xSet (s, x) z) - ls false = false -> - In z ls -> - goodBlinding_c_Prop k_X xinds f xSet s x z. - - induction ls; intuition; simpl in *. - - unfold goodBlinding_q_Prop. - intuition. - - intuition; subst. - unfold setLet in *. - unfold goodBlinding_c_Prop; intuition. - rewrite H1 in H. - - case_eq (badBlinding_ez xinds f xSet (F_P k_X x) z0); intuition. - rewrite H2 in H. - rewrite fold_left_bad_true in H. - discriminate. - - eapply badBlinding_ez_Prop_correct; intuition. - - unfold setLet in *. - - case_eq ( @arrayLookup Blist nat (@list_EqDec bool bool_EqDec) f - (@app bool s (natToBlist a))); intuition; - rewrite H1 in H. - - case_eq ((badBlinding_ez xinds f xSet (F_P k_X x) n)); intuition. - rewrite H3 in H. - rewrite fold_left_bad_true in H. - discriminate. - - rewrite H3 in H. - - eapply IHls. - eapply H. - trivial. - - eapply IHls. - eapply H. - trivial. - Qed. - - Theorem goodBlinding_q_Prop_correct : - forall k_X xinds f xSet s x, - badBlinding_q k_X xinds f xSet (s, x) = false -> - goodBlinding_q_Prop k_X xinds f xSet s x. - - unfold badBlinding_q, goodBlinding_q_Prop. - intuition. - eapply goodBlinding_q_Prop_correct_h; eauto. - - Qed. - - Definition natRandFunc := (@randomFunc Blist nat (RndNat (2^lambda)) _ ). - - Definition G10_3 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, f] <-$2 (Initialize_10_2 k_X k_I db q sigmas xSet) _ _ natRandFunc nil; - [edb, ls] <-2 z; - b <-$ A2 s_A edb (snd (split ls)); - ret (b, badBlinding k_I k_X db f xSet q). - - (* Step 10.4 : use the bookkeeping and only unblind when the z values match *) - Definition G10_4 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, f] <-$2 (Initialize_11 k_X k_I db q sigmas xSet) _ _ natRandFunc nil; - [edb, ls] <-2 z; - b <-$ A2 s_A edb (snd (split ls)); - ret (b, badBlinding k_I k_X db f xSet q). - - - Hypothesis natGroupOp_exp : - forall a b c, - a ^ (b * c)%group = a ^ (b * c)%nat. - - Hypothesis MaxMatchingIDs_correct : - forall db q s_A, - In (db, q, s_A) (getSupport A1) -> - forall w, - (length (lookupInds db w) <= MaxMatchingIDs)%nat. - - Theorem compMap_support_unary : - forall (A B : Set)(eqdb : EqDec B)(P : B -> Prop)(lsa : list A)(c : A -> Comp B)(lsb : list B), - In lsb (getSupport (compMap _ c lsa)) -> - (forall a b, In a lsa -> In b (getSupport (c a)) -> P b) -> - forall b, In b lsb -> P b. - - Local Opaque getSupport. - - induction lsa; intuition; simpl in *; - repeat simp_in_support. - simpl in *. - intuition. - - simpl in *. - intuition; subst. - eapply H0; intuition. - eapply IHlsa; eauto. - - Qed. - - Theorem natRandFunc_badBlinding_persists : - (forall q f1 f2 xSet db k_X k_I a b, - badBlinding k_I k_X db f1 xSet q = true -> - In (b, f2) (getSupport (natRandFunc f1 a)) -> - badBlinding k_I k_X db f2 xSet q = true). - - intuition. - unfold natRandFunc, randomFunc in *. - case_eq (arrayLookup (list_EqDec bool_EqDec) f1 a); intuition; - unfold Blist in *; - rewrite H2 in H1; - repeat simp_in_support. - trivial. - - - Theorem badBlinding_ez_f_app : - forall xinds f1 xSet y n a x, - badBlinding_ez xinds f1 xSet y n = true -> - badBlinding_ez xinds ((a, x) :: f1) xSet y n = true. - - induction xinds; - - intuition. - unfold badBlinding_ez in *. - unfold setLet in *. - unfold allPairs_2 in *. - simpl in *. - - Theorem fold_left_bool_app_true : - forall (A : Set)(ls1 ls2 : list A) f, - fold_left (fun (b : bool) a => if b then true else (f a)) (ls1 ++ ls2) false = true -> - fold_left (fun (b : bool) a => if b then true else (f a)) ls1 false = true \/ - fold_left (fun (b : bool) a => if b then true else (f a)) ls2 false = true. - - induction ls1; intuition; simpl in *. - destruct (f a). - left. - eapply fold_left_bad_true. - - eapply IHls1. - trivial. - - Qed. - - Theorem fold_left_bool_app_true_l : - forall (A : Set)(ls1 ls2 : list A) f, - fold_left (fun (b : bool) a => if b then true else (f a)) ls1 false = true -> - fold_left (fun (b : bool) a => if b then true else (f a)) (ls1 ++ ls2) false = true. - - induction ls1; intuition; simpl in *. - destruct (f a). - eapply fold_left_bad_true. - eauto. - - Qed. - - Theorem fold_left_bool_app_true_r : - forall (A : Set)(ls1 ls2 : list A) f, - fold_left (fun (b : bool) a => if b then true else (f a)) ls2 false = true -> - fold_left (fun (b : bool) a => if b then true else (f a)) (ls1 ++ ls2) false = true. - - induction ls1; intuition; simpl in *. - destruct (f a). - eapply fold_left_bad_true. - eauto. - - Qed. - - specialize (IHxinds f1). - remember (split f1) as z. - destruct z. - simpl in *. - destruct (eq_nat_dec x n); subst. - - eapply fold_left_bool_app_true in H; intuition. - eapply fold_left_bool_app_true_l . - trivial. - - eapply fold_left_bool_app_true_r . - eapply IHxinds; eauto. - - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (y * n * a * inverse x)) xSet). - eapply fold_left_bad_true. - - eapply fold_left_bool_app_true in H; intuition. - eapply fold_left_bool_app_true_l . - trivial. - eapply fold_left_bool_app_true_r . - eapply IHxinds; eauto. - - Qed. - - Theorem badBlinding_c_f_app : - forall q f1 xSet k_X xinds a x c, - arrayLookup _ f1 a = None -> - badBlinding_c k_X xinds f1 xSet q c = true -> - badBlinding_c k_X xinds ((a, x) :: f1) xSet q c = true. - - intuition. - unfold badBlinding_c in *. - destruct q. - unfold setLet in *. - - simpl. - case_eq (@eqb Blist (@list_EqDec bool bool_EqDec) - (@app bool k (natToBlist c)) a); intuition. - rewrite eqb_leibniz in H2. - subst. - rewrite H in H1. - discriminate. - - case_eq (@arrayLookup Blist nat (@list_EqDec bool bool_EqDec) f1 - (@app bool k (natToBlist c)) ); intuition. - rewrite H3 in H1. - eapply badBlinding_ez_f_app ; intuition. - - rewrite H3 in H1. - discriminate. - - Qed. - - Theorem badBlinding_q_f_app_h : - forall ls q f1 xSet k_X xinds a x , - arrayLookup _ f1 a = None -> - fold_left - (fun (b : bool) (z : nat) => - if b then true else badBlinding_c k_X xinds f1 xSet q z) - ls false = true -> - fold_left - (fun (b : bool) (z : nat) => - if b then true else badBlinding_c k_X xinds ((a, x) :: f1) xSet q z) - ls false = true. - - induction ls; intuition; simpl in *. - - case_eq ((badBlinding_c k_X xinds f1 xSet q a)); intuition. - erewrite badBlinding_c_f_app. - eapply fold_left_bad_true. - trivial. - eauto. - - rewrite H2 in H1. - destruct ((badBlinding_c k_X xinds ((a0, x) :: f1) xSet q a)). - eapply fold_left_bad_true. - eapply IHls. - trivial. - trivial. - Qed. - - Theorem badBlinding_q_f_app : - forall q f1 xSet k_X xinds a x , - arrayLookup _ f1 a = None -> - badBlinding_q k_X xinds f1 xSet q = true -> - badBlinding_q k_X xinds ((a, x) :: f1) xSet q = true. - - intuition. - unfold badBlinding_q in *. - eapply badBlinding_q_f_app_h. - trivial. - trivial. - Qed. - - Theorem badBlinding_f_app : - forall q f1 xSet db k_X k_I a x , - arrayLookup _ f1 a = None -> - badBlinding k_I k_X db f1 xSet q = true -> - badBlinding k_I k_X db ((a, x) :: f1) xSet q = true. - - induction q; intuition. - unfold badBlinding. - unfold badBlinding in H1. - unfold setLet in *. - simpl in *. - - case_eq ((badBlinding_q k_X - (foreach (x0 in fst (split db))F_P k_I (Vector.to_list x0)) f1 - xSet a)); intuition. - erewrite badBlinding_q_f_app. - eapply fold_left_bad_true. - trivial. - trivial. - rewrite H2 in H1. - - case_eq (badBlinding_q k_X - (foreach (x0 in fst (split db))F_P k_I (Vector.to_list x0)) - ((a0, x) :: f1) xSet a); intuition. - eapply fold_left_bad_true. - unfold badBlinding in *. - unfold setLet in *. - eapply IHq. - trivial. - eauto. - Qed. - - eapply badBlinding_f_app. - trivial. - trivial. - Qed. - - Theorem natRandFunc_badBlinding_persists_oc : - (forall (X : Set) (c : OracleComp _ _ X) (a : X) f1 f2 q xSet db k_X k_I, - badBlinding k_I k_X db f1 xSet q = true -> - In (a, f2) (getSupport (c _ _ natRandFunc f1)) -> - badBlinding k_I k_X db f2 xSet q = true) . - - intuition. - eapply (@oc_comp_invariant _ _ _ c _ (fun s => badBlinding k_I k_X db s xSet q = true) _ natRandFunc); intuition. - - eapply natRandFunc_badBlinding_persists; - eauto. - eauto. - eauto. - - Qed. - - Definition goodBlinding_Prop k_I k_X (db : DB) f xSet qs := - xinds <- (foreach (x in fst (split db))F_P k_I (Vector.to_list x)); - forall s x, In (s, x) qs -> - goodBlinding_q_Prop k_X xinds f xSet s x. - - Theorem goodBlinding_Prop_correct : - forall qs k_I k_X db f xSet, - badBlinding k_I k_X db f xSet qs = false -> - goodBlinding_Prop k_I k_X db f xSet qs. - - induction qs; - intuition; - unfold badBlinding, goodBlinding_Prop in *; - unfold setLet in *; - intuition. - - simpl in *. - intuition. - - simpl in *. - intuition. - subst. - - case_eq (badBlinding_q k_X - (foreach (x0 in fst (split db))F_P k_I (Vector.to_list x0)) f - xSet (s, x)); intuition; - rewrite H in H0. - rewrite fold_left_bad_true in H0. - discriminate. - - eapply goodBlinding_q_Prop_correct; intuition. - - case_eq (badBlinding_q k_X - (foreach (x in fst (split db))F_P k_I (Vector.to_list x)) f xSet - a); intuition; - rewrite H in H0. - rewrite fold_left_bad_true in H0. - discriminate. - - eapply IHqs. - eauto. - trivial. - Qed. - - - Theorem oc_compMap_wf : - forall (A B C D : Set)(eqdd : EqDec D)(ls : list A)(c : A -> OracleComp B C D)(S : Set)(eqds : EqDec S)(inv : S -> Prop)(o : S -> B -> Comp (C * S))(s : S), - (forall a, In a ls -> well_formed_oc (c a)) -> - (forall s a, inv s -> well_formed_comp (o s a)) -> - inv s -> - (forall s s' b c, inv s -> In (c, s') (getSupport (o s b)) -> inv s') -> - well_formed_comp ((oc_compMap _ c ls) _ _ o s). - - induction ls; intuition; simpl in *; - wftac. - - eapply oc_comp_wf_inv. - eauto. - intuition. - eapply H0. - eapply H3. - intuition. - eapply H2; intuition. - eapply H3. - eauto. - trivial. - - eapply IHls; - intuition. - eapply H0. - eapply H4. - eapply oc_comp_invariant. - intuition. - eapply H2. - eapply H5. - eauto. - eapply H1. - eauto. - eapply H2. - eapply H4. - eauto. - Qed. - - - Theorem oc_compMap_spec_bad: - forall (D E S A1 A2 : Set) P1 P2 (inv : S -> Prop)(bad : S -> bool) (B C : Set) (eqdd : EqDec D)(eqde : EqDec E) (f1 : A1 -> OracleComp B C D)(f2 : A2 -> OracleComp B C E) - (ls1 : list A1)(ls2 : list A2)(eqds : EqDec S) (o : S -> B -> Comp (C * S)) - (s : S) , - list_pred P1 ls1 ls2 -> - bad s = false -> - inv s -> - (forall a, In a ls1 -> well_formed_oc (f1 a)) -> - (forall a, In a ls2 -> well_formed_oc (f2 a)) -> - (forall s a, inv s -> well_formed_comp (o s a)) -> - (forall s s' a b , - inv s -> In (b, s') (getSupport (o s a)) -> inv s') -> - (forall (s0 : S) (a1 : A1)(a2 : A2), - bad s0 = false -> - inv s0 -> - P1 a1 a2 -> - In a1 ls1 -> - In a2 ls2 -> - comp_spec (fun a b => bad (snd a) = bad (snd b) /\ - (bad (snd a) = false -> ((P2) (fst a) (fst b) /\ snd a = snd b))) - ((f1 a1) S eqds o s0) ((f2 a2) S eqds o s0)) -> - (forall s1 s2 a b, - bad s1 = true -> - In (b, s2) (getSupport (o s1 a)) -> - bad s2 = true) -> - comp_spec (fun a b => bad (snd a) = bad (snd b) /\ - (bad (snd a) = false -> (list_pred (P2) (fst a) (fst b) /\ snd a = snd b))) - ((oc_compMap _ f1 ls1) S eqds o s) - ((oc_compMap _ f2 ls2) S eqds o s). - - induction ls1; inversion 1; intuition; simpl in *. - - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - econstructor. - - subst. - comp_skip. - comp_simp. - simpl in *; subst. - case_eq (bad b); intuition. - comp_irr_l. - - eapply oc_compMap_wf; intuition. - eapply H9; intuition. - eauto. - eapply oc_comp_invariant. - eauto. - eauto. - eauto. - eapply H10; eauto. - - comp_irr_r. - eapply oc_compMap_wf; intuition. - eapply H9; intuition. - eauto. - eapply oc_comp_invariant. - eauto. - eauto. - eauto. - eapply H10; eauto. - - comp_simp. - - assert (bad s1 = true). - eapply oc_comp_invariant_f. - intuition. - eapply H12. - eauto. - eauto. - eauto. - eapply H15. - - eapply comp_spec_ret; intros. - simpl. - split. - - rewrite H17. - assert (bad s2 = true). - eapply oc_comp_invariant_f. - intuition. - eapply H12. - eapply H19. - eauto. - rewrite H13 in H3. - eapply H3. - eauto. - rewrite H18. - trivial. - - intros. - exfalso. - congruence. - - subst. - comp_skip. - eapply IHls1; intuition. - eapply oc_comp_invariant; - eauto. - - comp_simp. - simpl in *. - eapply comp_spec_ret; intuition. - subst. - simpl. - econstructor; intuition. - Qed. - - - Theorem oc_compMap_support_inv : - forall (S : Set) (inv : S -> Prop)(A B C D : Set) (eqdd : EqDec D) (c : A -> OracleComp B C D) (ls : list A) (eqds : EqDec S)(o : S -> B -> Comp (C * S)) s x y, - (forall x y a d, inv x -> In (d, y) (getSupport ((c a) _ _ o x)) -> inv y) -> - inv s -> - In (x, y) (getSupport ((oc_compMap _ c ls) _ _ o s)) -> - inv y. - - induction ls; intuition. simpl in *. - repeat simp_in_support. - trivial. - - Local Opaque getSupport. - simpl in *. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - eapply IHls. - intuition. - eapply H. - eapply H1. - eauto. - eapply H. - eapply H0. - eauto. - eauto. - Qed. - - Theorem oc_compMap_support : - forall (S A D : Set)(R : S -> A -> D -> Prop)(B C : Set)(eqdd : EqDec D)(c : A -> OracleComp B C D)(ls : list A)(eqds : EqDec S)(o : S -> B -> Comp (C * S)) s x y , - (forall s s' a d, In a ls -> In (d, s') (getSupport ((c a) _ _ o s)) -> R s' a d) -> - (forall s s' a d a' d', R s a d -> In (d', s') (getSupport ((c a') _ _ o s)) -> R s' a d) -> - In (x, y) (getSupport ((oc_compMap _ c ls) _ _ o s)) -> - list_pred (R y) ls x. - - induction ls; intuition. - simpl in *. - repeat simp_in_support. - econstructor. - - simpl in *. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - econstructor. - eapply (@oc_compMap_support_inv _ (fun s => R s a d) _ _ _ _ _ c). - intros. - eapply H0. - eapply H1. - eapply H4. - eapply H. - eauto. - eauto. - eauto. - - eapply IHls. - intuition. - eapply H; eauto. - intuition. - eauto. - Qed. - - Theorem list_pred_In_exists : - forall (A B : Set) (a : A) lsa (lsb : list B)(P : A -> B -> Prop), - list_pred P lsa lsb -> - In a lsa -> - exists b, - In b lsb /\ P a b. - - induction lsa; inversion 1; intuition; simpl in *; subst; - intuition. - subst. - econstructor. - intuition. - inversion H; clear H; subst. - edestruct IHlsa. - eauto. - trivial. - intuition. - econstructor. - split. - right. - eapply H1. - trivial. - - Qed. - - Theorem In_allPairs_2 : - forall (A B : Type)(lsa : list A)(lsb : list B) a b, - In a lsa -> - In b lsb -> - In (a, b) (allPairs_2 lsa lsb). - - induction lsa; intuition; simpl in *. - intuition. - subst. - unfold allPairs_2. - simpl. - eapply in_or_app. - left. - unfold allPairs_1. - eapply in_map_iff. - econstructor. - intuition. - - unfold allPairs_2. - simpl. - eapply in_or_app. - right. - eapply IHlsa; trivial. - - Qed. - - Theorem randomFunc_In_inv : - forall (A B C : Set)(eqda : EqDec A)(eqdb : EqDec B)(c : OracleComp A B C) s s' x (b : B) r, - In b (snd (split s)) -> - In (x, s') (getSupport (c _ _ (randomFunc r _) s)) -> - In b (snd (split s')). - - intuition. - eapply (@oc_comp_invariant _ _ _ c _ (fun s => In b (snd (split s))) _ (randomFunc r _)); intuition. - unfold randomFunc in *. - case_eq (arrayLookup eqda c0 d); intuition; - rewrite H3 in H1; - repeat simp_in_support. - trivial. - simpl. - remember (split c0) as z. - destruct z. - simpl in *. - intuition. - eauto. - eauto. - Qed. - - Theorem GenTrans_10_1_11_spec : - forall a1 xSet k_I k_X q qs a2 b3 s0 db, - (length b3 <= MaxMatchingIDs)%nat -> - (forall a b c, In (a, (b, c)) b3 -> In c (snd (split s0)) /\ In b (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a)) = F_P k_I (Vector.to_list b) * inverse c )%group -> - badBlinding k_I k_X db s0 xSet qs = false -> - In q qs -> - comp_spec (fun a b => ( - badBlinding k_I k_X db (snd a) xSet qs = - badBlinding k_I k_X db (snd b) xSet qs) /\ - (badBlinding k_I k_X db (snd a) xSet qs = false -> - a = b)) - ((GenTrans_10_1 (a1, xSet) k_X (q, (a2, b3))) - (list (Blist * nat)) - (list_EqDec (pair_EqDec (list_EqDec bool_EqDec) nat_EqDec)) natRandFunc s0) - ((GenTrans_11 (a1, xSet) k_X (q, (a2, b3))) (list (Blist * nat)) - (list_EqDec (pair_EqDec (list_EqDec bool_EqDec) nat_EqDec)) natRandFunc s0). - - intuition. - intros. - unfold GenTrans_10_1, GenTrans_11. - comp_simp. - simpl. - comp_skip. - - comp_simp. - eapply comp_spec_ret; intuition. - simpl in *. - - assert ((ServerSearch_10_1 xSet a0 b3) = - (ServerSearch_11 xSet a0 b3)). - unfold ServerSearch_10_1, ServerSearch_11. - - eapply map_ext_in. - intuition. - unfold ServerSearchLoop_10_1, ServerSearchLoop_11. - - destruct a. - destruct p0. - destruct p1. - destruct p0. - remember (splitVector lambda lambda b0) as z. - destruct z. - - destruct (eq_nat_dec n0 n1); trivial. - destruct (in_dec (EqDec_dec nat_EqDec) (n ^ bvectorToNat t0) xSet); trivial. - exfalso. - - eapply goodBlinding_Prop_correct in H7. - eapply n2. - - eapply (@oc_compMap_support _ _ _ (fun s (a : nat)(b : nat * nat) => arrayLookup _ s (k ++ a) = Some (snd b) /\ fst b = g ^ (F_P k_X k0 * (snd b)))) in H0. - - pose proof H8. - eapply in_combine_l in H8. - edestruct list_pred_In_exists. - eapply list_pred_symm in H0. - eapply H0. - eapply H8. - intuition. - simpl in *; subst. - - symmetry. - eapply H7. - eauto. - eapply allNatsLt_lt_if. - eapply lt_le_trans. - eapply allNatsLt_lt. - eapply H11. - trivial. - trivial. - - eapply In_allPairs_2. - eapply in_map_iff. - exists b1. - intuition. - eapply H2. - eapply in_combine_r. - eauto. - - eapply (@ oc_compMap_support_inv (list (Blist * nat)) (fun s => In n1 (snd (split s))) - nat Blist nat (nat * nat) ); intros. - - eapply randomFunc_In_inv. - eauto. - eauto. - eapply H2. - eapply in_combine_r. - eauto. - eapply H6. - - simpl in *. - subst. - assert ( bvectorToNat t0 = F_P k_I (Vector.to_list b1) * inverse n1)%group. - assert (t0 = snd (splitVector lambda lambda b0)). - rewrite <- Heqz. - trivial. - subst. - eapply H2. - eapply in_combine_r. - eauto. - rewrite H12 in i. - - rewrite natGroupOp_exp in i. - rewrite groupExp_mult in i. - rewrite mult_assoc in i. - trivial. - - intuition. - simpl in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - simpl. - unfold natRandFunc, randomFunc in *. - case_eq ( @arrayLookup Blist nat (@list_EqDec bool bool_EqDec) s - (@app bool k (natToBlist a))); intuition; - rewrite H10 in H11. - repeat simp_in_support. - trivial. - repeat simp_in_support. - simpl. - case_eq (eqb (k ++ a) (k ++ a)); intuition. - rewrite eqb_refl in H11. - discriminate. - - simpl in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - simpl. - trivial. - - intuition. - simpl in *. - subst. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - unfold natRandFunc, randomFunc in *. - case_eq (@arrayLookup Blist nat (@list_EqDec bool bool_EqDec) s - (@app bool k (natToBlist a'))); intuition; - rewrite H9 in H10; - repeat simp_in_support. - trivial. - simpl. - case_eq (eqb (k ++ a) (k ++ a')); intuition. - rewrite eqb_leibniz in H10. - - apply app_inv_head in H10. - rewrite H10 in H11. - unfold Blist in H9. - rewrite H9 in H11. - discriminate. - - rewrite H8. - trivial. - Qed. - - Theorem list_pred_unary_in : - forall (A : Set) P (ls : list A), - list_pred P ls ls -> - (forall a, In a ls -> P a a). - - induction ls; intuition; simpl in *; - intuition. - subst. - inversion H; clear H; subst. - trivial. - inversion H; clear H; subst. - eapply IHls; eauto. - - Qed. - - - Theorem list_pred_in_all : - forall (A B : Set)(P1 : A -> Prop)(P2 : B -> Prop)(lsa : list A)(lsb : list B), - list_pred (fun a b => P1 a -> P2 b) lsa lsb -> - (forall a, In a lsa -> P1 a) -> - forall b, In b lsb -> P2 b. - - induction 1; intuition; simpl in *. - intuition. - intuition; subst. - eauto. - - Qed. - - Theorem oc_compMap_length : - forall (A B C D : Set)(eqdd : EqDec D)(ls : list A)(c : A -> OracleComp B C D)(S : Set)(eqds : EqDec S)(o : S -> B -> Comp (C * S))(s : S) x y, - In (x, y) (getSupport ((oc_compMap _ c ls) _ _ o s)) -> - length x = length ls. - - induction ls; intuition; simpl in *; - repeat simp_in_support; - simpl. - trivial. - destruct x0. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - simpl. - f_equal. - eauto. - - Qed. - - Theorem randomFunc_In_inv_once: - forall (A B : Set)(eqda : EqDec A) (s s' : list (A * B)) a b b' r, - In b (snd (split s)) -> - In (b', s') - (getSupport (randomFunc r _ s a)) -> In b (snd (split s')). - - - intuition. - unfold randomFunc in *. - destruct (arrayLookup eqda s a); - repeat simp_in_support. - trivial. - simpl. - remember (split s) as z. - destruct z. - simpl in *. - intuition. - - Qed. - - Hypothesis nat_Bvector_inverse : - forall n, n < 2 ^ lambda -> - bvectorToNat - (natToBvector n) = n. - - Hypothesis group_lt_lambda : - forall n1 n2, - (n1 * n2)%group < 2 ^ lambda. - - Theorem Init_10_2_11_eq_until_bad : - forall k_X k_I db q sigmas xSet s, - (forall z, length (lookupInds db z) <= MaxMatchingIDs)%nat -> - (forall z, In z sigmas -> length z <= MaxMatchingIDs)%nat -> - length (toW db) = length sigmas -> - (forall k l1, In (k, l1) (combine (toW db) sigmas) -> - forall x, In x l1 -> - x < length (lookupInds db k)) -> - comp_spec - (fun a b => (badBlinding k_I k_X db (snd a) xSet q = badBlinding k_I k_X db (snd b) xSet q) /\ - (badBlinding k_I k_X db (snd a) xSet q = false -> fst a = fst b)) - ((Initialize_10_2 k_X k_I db q sigmas xSet) _ _ natRandFunc s) - ((Initialize_11 k_X k_I db q sigmas xSet) _ _ natRandFunc s). - - intuition. - - unfold Initialize_10_2, Initialize_11. - simpl. - - assert TSet. - assert (TSet * TSetKey). - eapply comp_base_exists. - eapply TSetSetup. - econstructor. - intuition. - - comp_skip. - - remember (split a0) as z. - destruct z. - - simpl. - inline_first. - comp_skip. - - simpl. - comp_simp. - simpl. - inline_first. - comp_skip. - comp_simp. - - case_eq (badBlinding k_I k_X db b xSet q); intuition. - comp_irr_l. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - eapply oc_compMap_wf; intuition. - - Theorem oc_compMap_wfoc : - forall (A B C D : Set)(eqdd : EqDec D)(c : A -> OracleComp B C D)(ls : list A), - (forall a, In a ls -> well_formed_oc (c a)) -> - well_formed_oc (oc_compMap _ c ls). - - induction ls; intuition; simpl in *. - econstructor. - wftac. - - econstructor. - eapply H. - intuition. - - intuition. - econstructor. - eapply IHls. - intuition. - intuition. - econstructor. - wftac. - Qed. - - Theorem GenTrans_10_1_wf : - forall x y z, - well_formed_oc (GenTrans_10_1 x y z). - - intuition. - unfold GenTrans_10_1. - comp_simp. - econstructor. - eapply oc_compMap_wfoc; intuition. - econstructor. - econstructor. - intuition. - econstructor. - wftac. - - intuition. - econstructor. - wftac. - Qed. - - eapply GenTrans_10_1_wf. - eapply randomFunc_wf. - eapply well_formed_RndNat. - assert (nz (2 ^ lambda)). - eapply expnat_nz; intuition. - eapply H. - assert ((fun x => True) b). - intuition. - eapply H13. - intuition. - - comp_irr_r. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - eapply pair_EqDec; intuition. - eapply list_EqDec; intuition. - - eapply oc_compMap_wf; intuition. - Theorem GenTrans_11_wf : - forall x y z, - well_formed_oc (GenTrans_11 x y z). - - intuition. - unfold GenTrans_11. - comp_simp. - econstructor. - eapply oc_compMap_wfoc; intuition. - econstructor. - econstructor. - intuition. - econstructor. - wftac. - - intuition. - econstructor. - wftac. - Qed. - eapply GenTrans_11_wf. - eapply randomFunc_wf. - eapply well_formed_RndNat. - assert (nz (2 ^ lambda)). - eapply expnat_nz; intuition. - eapply H. - assert ((fun x => True) b). - intuition. - eapply H14. - intuition. - - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - - erewrite natRandFunc_badBlinding_persists_oc. - symmetry. - erewrite natRandFunc_badBlinding_persists_oc. - trivial. - eauto. - eapply H14. - eauto. - eapply H13. - simpl in *. - assert (badBlinding k_I k_X db l2 xSet q = true). - eapply natRandFunc_badBlinding_persists_oc. - eauto. - eauto. - congruence. - - assert (forall p, In p b1 -> forall a x c, In (a, (x, c)) (snd p) -> In c (snd (split b)) /\ In x (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a)) = F_P k_I (Vector.to_list x) * inverse c )%group. - - intros. - - eapply (@oc_compMap_support _ _ _ (fun s a b => forall a x c, In (a, (x, c)) (fst b) -> - In c (snd (split s)) /\ In x (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a)) = F_P k_I (Vector.to_list x) * inverse c)) in H6. - - eapply list_pred_symm in H6. - - assert (list_pred (fun - (b0 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) - (_ : _) => - forall (a1 : Vector.t bool (lambda + lambda)) - (x : Identifier lambda) (c : nat), - In (a1, (x, c)) b0 -> - In c (snd (split b)) /\ - In x (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a1)) = - F_P k_I (Vector.to_list x) * inverse c) - (fst (split a0)) (combine (toW db) sigmas)). - eapply list_pred_fst_split_irr. - eapply H6. - clear H6. - - - eapply (@list_pred_impl_unary _ _ _ - ((fun - (b0 _ : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) => - forall (a1 : Vector.t bool (lambda + lambda)) - (x : Identifier lambda) (c : nat), - In (a1, (x, c)) b0 -> - In c (snd (split b)) /\ - In x (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a1)) = - F_P k_I (Vector.to_list x) * inverse c)) - ) - in H15. - - eapply list_pred_unary_in in H15. - eapply H15. - eapply H14. - eapply (@compMap_support _ _ (fun a b => snd b = nil \/ In (snd b) l)) in H10. - eapply list_pred_symm in H10. - assert ( list_pred - (fun - (b2 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) (_ : Keyword) => - b2 = nil \/ In b2 l) (snd (split b1)) (fst (split q))). - eapply list_pred_snd_split_irr. - eauto. - clear H10. - eapply (@list_pred_impl_unary _ _ _ (fun a _ => a = nil \/ In a l)) in H6. - rewrite <- Heqz. - simpl. - eapply list_pred_unary_in in H6. - intuition; eauto. - rewrite H10 in H14. - simpl in *. - intuition. - - Theorem snd_split_in : - forall (A B : Type)(ls : list (A * B))(p : A * B), - In p ls -> - In (snd p) (snd (split ls)). - - induction ls; intuition; simpl in *. - intuition; subst. - simpl. - remember (split ls) as z. - destruct z. - simpl. - intuition. - - remember (split ls) as z. - destruct z. - simpl. - right. - eapply IHls. - trivial. - Qed. - - eapply snd_split_in. - trivial. - - intuition. - intuition. - repeat simp_in_support. - simpl. - unfold arrayLookupList. - case_eq ( @arrayLookup Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (@list_EqDec bool bool_EqDec) - (@combine Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (toW db) l) a2); intuition. - - Theorem arrayLookup_combine_Some_In : - forall (A B : Set)(eqd : EqDec A) (lsa : list A) (lsb : list B) a b, - arrayLookup _ (combine lsa lsb) a = Some b -> - In b lsb. - - induction lsa; intuition; simpl in *. - discriminate. - destruct lsb; - simpl in *. - discriminate. - - destruct (eqb a0 a ). - inversion H; clear H; subst. - intuition. - right. - eauto. - - Qed. - - eapply arrayLookup_combine_Some_In in H5. - intuition. - - intuition. - intros. - unfold Initialize_wLoop_11 in *. - destruct a2. - - Ltac simp_in_support' := - unfold setLet in *; - match goal with - | [H : In _ (getSupport (Bind _ _)) |- _ ] => - apply getSupport_Bind_In in H; destruct_exists - | [H : In _ (getSupport (if ?t then _ else _)) |- _ ] => let x := fresh "x" in remember t as x; destruct x - | [H : In _ (getSupport (ret _)) |- _ ] => apply getSupport_In_Ret in H; try pairInv; subst - (* | [H : false = inRange _ _ |- _] => symmetry in H *) - | [H : true = negb ?t |- _ ] => let x := fresh "x" in remember t as x; destruct x; simpl in H; try discriminate - | [H : _ /\ _ |- _ ] => destruct H - end. - - unfold setLet in *. - simpl in *. - repeat simp_in_support'. - destruct x1. - repeat simp_in_support'. - simpl in *. - - eapply (@oc_compMap_support _ _ _ (fun (l3 : list (Blist * nat)) (a : Identifier lambda * nat) (b : Vector.t bool (lambda + lambda) * (Identifier lambda * nat)) => - In (fst a) (fst (split db)) -> - let a3 := fst b in let (x0, c0) := snd b in - (In c0 (snd (split l3)) /\ - In x0 (fst (split db)) /\ - bvectorToNat (snd (splitVector lambda lambda a3)) = - F_P k_I (Vector.to_list x0) * inverse c0)) - - ) in H18. - - - Theorem list_pred_impl_all : - forall (A B : Set) P1 (P2 : B -> Prop) (lsa : list A)(lsb : list B), - list_pred (fun a b => P1 a -> P2 b) lsa lsb -> - forall b, - (forall a, In a lsa -> P1 a) -> - In b lsb -> P2 b. - - induction 1; intuition; simpl in *. - intuition. - - intuition; subst. - eapply H. - eapply X. - intuition. - - Qed. - - specialize (@list_pred_impl_all _ _ _ _ _ _ H18 (a3, (x0, c0))); intros. - simpl in *. - eapply H19. - intuition. - destruct a2. - simpl. - eapply in_combine_l in H. - - Theorem in_permute_if : - forall (A : Set)(sigma : list nat)(ls : list A)(a : A), - In a (permute ls sigma) -> - (forall n, In n sigma -> n < length ls) -> - In a ls. - - induction sigma; intuition; simpl in *; - intuition. - - case_eq (nth_error ls a); intuition; - rewrite H1 in H. - simpl in *; intuition. - subst. - - Theorem nth_error_Some_In : - forall (A : Set)(ls : list A) n (a : A), - nth_error ls n = Some a -> - In a ls. - - induction ls; destruct n; intuition; simpl in *. - discriminate. - discriminate. - inversion H; clear H; subst. - intuition. - - right. - eapply IHls. - eauto. - - Qed. - - eapply nth_error_Some_In. - eauto. - simpl in *. - intuition. - - Qed. - - Theorem In_lookupInds_if_h : - forall db w i z, - In i - (fold_left - (fun (ls : list (Identifier lambda)) - (p0 : Identifier lambda * list Keyword) => - if in_dec (EqDec_dec (list_EqDec bool_EqDec)) w (snd p0) - then fst p0 :: ls - else ls) db z) -> - In i ((fst (split db)) ++ z). - - induction db; intuition; simpl in *. - - remember (split db) as y. - destruct y. - simpl in *. - - specialize (@IHdb _ _ _ H0). - eapply in_app_or in IHdb. - intuition. - - destruct (@in_dec (list bool) - (@EqDec_dec (list bool) (@list_EqDec bool bool_EqDec)) w b). - simpl in *. - intuition. - right. - eapply in_or_app. - intuition. - - Qed. - - - Theorem In_lookupInds_if : - forall db w i, - In i (lookupInds db w) -> - In i (fst (split db)). - - intuition. - specialize (@In_lookupInds_if_h _ _ _ _ H); intuition. - apply in_app_or in H0. - intuition. - simpl in *. - intuition. - - Qed. - - eapply In_lookupInds_if. - eapply in_permute_if. - eapply H. - intuition. - eapply H3. - eauto. - trivial. - - trivial. - - intuition. - intuition. - unfold Initialize_indLoop_11 in *. - simpl in *. - repeat simp_in_support'. - destruct x1. - repeat simp_in_support'. - simpl. - intuition. - - Theorem arrayLookup_Some_In : - forall (A B : Set)(eqd : EqDec A) (lsb : list (A * B)) a b, - arrayLookup _ lsb a = Some b -> - In b (snd (split lsb)). - - induction lsb; intuition; simpl in *. - discriminate. - remember (split lsb) as z. - destruct z. - simpl. - - destruct (eqb a a0 ). - inversion H; clear H; subst. - intuition. - right. - eauto. - - Qed. - - Theorem valueInRandomFunc : - forall (A B : Set)(eqda : EqDec A)(a : A)(v : B) f f' r, - In (v, f') (getSupport (randomFunc r _ f a)) -> - In v (snd (split f')). - - intuition. - unfold randomFunc in *. - case_eq (arrayLookup eqda f a); intuition; - rewrite H0 in H; - repeat simp_in_support. - eapply arrayLookup_Some_In. - eauto. - simpl. - remember (split f) as z. - destruct z. - simpl. - intuition. - - Qed. - - unfold natRandFunc in *. - eapply valueInRandomFunc. - eauto. - - rewrite splitVector_append. - simpl. - eapply nat_Bvector_inverse. - eapply group_lt_lambda. - - intuition. - simpl in *. - repeat simp_in_support'. - destruct x1. - repeat simp_in_support'. - intuition. - eapply randomFunc_In_inv_once; - eauto. - - intuition. - - eapply randomFunc_In_inv. - eapply H. - eapply H17. - eapply H16. - eapply H. - eapply H17. - eapply H. - trivial. - - comp_skip. - - eapply (@oc_compMap_spec_bad _ _ _ _ _ _ _ (fun s => forall x, In x b1 -> forall y, In y (snd x) -> In (snd (snd y)) (snd (split s))) - (fun b => badBlinding k_I k_X db b xSet q)). - eapply list_pred_eq. - trivial. - intuition. - destruct y. - simpl. - destruct p0. - simpl. - eapply H13. - eauto. - eauto. - intuition; subst. - - eapply GenTrans_10_1_wf. - - intuition. - eapply GenTrans_11_wf. - - intuition. - unfold natRandFunc. - eapply randomFunc_wf. - eapply well_formed_RndNat. - assert (nz (2 ^ lambda)). - eapply expnat_nz. - intuition. - inversion H; intuition. - - intuition. - - eapply randomFunc_In_inv_once. - eapply H. - eauto. - trivial. - eauto. - - intuition. - subst. - eapply comp_spec_consequence. - eapply GenTrans_10_1_11_spec. - - apply in_combine_r in H17. - eapply (@compMap_support _ _ (fun a b => snd b = arrayLookupList _ (combine (toW db) l) a)) in H10. - eapply list_pred_symm in H10. - eapply (@list_pred_impl_unary _ _ _ (fun a a => (snd a) = nil \/ In (snd a) l)) in H10. - - eapply list_pred_unary_in in H10. - intuition. - assert (snd (a2, b3) = nil). - eapply H16. - simpl in *. - subst. - simpl. - omega. - simpl in *. - - eapply (@oc_compMap_support _ _ _ (fun s a b => length (snd a) <= MaxMatchingIDs -> length (fst b) <= MaxMatchingIDs))%nat in H6. - eapply list_pred_symm in H6. - - eapply (@list_pred_fst_split_irr _ _ _ (fun - (b0 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) - (a : Keyword * list nat) => - (length (snd a) <= MaxMatchingIDs -> length b0 <= MaxMatchingIDs)%nat)) in H6. - - - - eapply list_pred_symm in H6. - eapply (@list_pred_snd_split_irr _ _ _ (fun - (a : list nat) - (b0 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) - => - (length a <= MaxMatchingIDs -> length b0 <= MaxMatchingIDs)%nat)) in H6. - - eapply (@list_pred_in_all _ _) in H6. - eapply H6. - intuition. - eapply H1. - assert (split (combine (toW db) sigmas) = ((toW db), sigmas)). - eapply combine_split. - trivial. - rewrite H19 in H10. - trivial. - rewrite <- Heqz. - trivial. - - intuition. - simpl in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - simpl. - - eapply oc_compMap_length in H22. - rewrite H22. - rewrite combine_length. - eapply Nat.min_le_iff. - right. - rewrite allNatsLt_length. - eapply H0. - - intuition. - trivial. - intuition. - - simpl in *; subst. - unfold arrayLookupList. - case_eq (@arrayLookup Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (@list_EqDec bool bool_EqDec) - (@combine Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (toW db) l) b4); intuition. - - - - apply arrayLookup_combine_Some_In in H. - intuition. - - intuition. - repeat simp_in_support. - trivial. - - intuition. - assert (c = (snd (snd (a3, (b2, c))))). - trivial. - rewrite H19. - clear H19. - eapply H15. - eapply in_combine_r. - eauto. - trivial. - eapply H13. - eapply in_combine_r. - eauto. - eauto. - eapply H13. - eapply in_combine_r. - eauto. - trivial. - - eauto. - eapply in_combine_l. - eauto. - - intuition. - simpl. - assert ((a3, (a4, b5)) = (fst b4)). - subst. - trivial. - eapply H20. - subst. - trivial. - - intuition. - eapply natRandFunc_badBlinding_persists; - eauto. - - comp_simp. - eapply comp_spec_ret; intuition. - simpl in *. - apply list_pred_eq_impl_eq in H18. - subst. - trivial. - - Qed. - - Theorem G10_4_eq_until_bad : - forall x, - evalDist G10_3 (x, false) == evalDist G10_4 (x, false). - - intuition. - unfold G10_3, G10_4. - do 4 (comp_skip; comp_simp). - - eapply comp_spec_impl_eq. - comp_skip. - - eapply Init_10_2_11_eq_until_bad. - eapply MaxMatchingIDs_correct; eauto. - unfold sampleSigmas_11 in *. - intuition. - - unfold setLet in *. - eapply (@compMap_support_unary _ _ _ (fun z => length z <= MaxMatchingIDs)%nat _ (fun w => RndPerm (length (lookupInds d w)))). - eapply H3. - intuition. - erewrite RndPerm_In_support_length. - eapply MaxMatchingIDs_correct. - eauto. - eauto. - trivial. - - symmetry. - erewrite compMap_length; eauto. - - intuition. - unfold sampleSigmas_11 in *. - unfold setLet in *. - eapply (@compMap_support _ _ (fun a b => forall x, In x b -> x < length (lookupInds d a))) in H3. - - Theorem list_pred_combine_equiv : - forall (A B : Set) P (lsa : list A)(lsb : list B), - list_pred P lsa lsb -> - forall x y, In (x, y) (combine lsa lsb) -> P x y. - - induction 1; intuition; simpl in *. - intuition. - - intuition. - pairInv. - trivial. - - Qed. - - eapply list_pred_combine_equiv in H3. - eapply H3; eauto. - trivial. - - intuition. - - Require Import Permutation. - - Theorem RndPerm_In_support_lt: - forall (n : nat) (ls : list nat), - In ls (getSupport (RndPerm n)) -> - forall x, In x ls -> x < n. - - intuition. - - eapply allNatsLt_lt. - eapply Permutation_in. - eapply Permutation_sym. - eapply RndPerm_In_support. - eauto. - trivial. - - Qed. - - eapply RndPerm_In_support_lt; - eauto. - - comp_simp. - intuition. - simpl in H9. - case_eq (badBlinding x1 x0 d b (XSetSetup_2 x0 x1 d x2) l); intuition. - comp_irr_l. - comp_irr_r. - simpl in H8. - rewrite <- H8. - rewrite H7. - eapply comp_spec_ret; intuition. - pairInv. - pairInv. - pairInv. - comp_skip. - simpl in H8. - rewrite <- H8. - rewrite H7. - eapply comp_spec_ret; intuition. - Qed. - - - Definition G10_3_bad := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - z <-$ (Initialize_10_2 k_X k_I db q sigmas xSet) _ _ natRandFunc nil; - ret (badBlinding k_I k_X db (snd z) xSet q). - - SearchAbout OracleComp. - - Theorem G10_3_bad_equiv : - Pr[x <-$ G10_3; ret snd x] == Pr[G10_3_bad]. - - unfold G10_3, G10_3_bad. - do 5 (comp_inline_l; comp_skip; comp_simp). - comp_inline_l. - comp_irr_l; wftac. - simpl. - intuition. - Qed. - - Require Import RndInList. - - Variable MaxKeywords : nat. - Hypothesis MaxKeywords_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length (toW db) <= MaxKeywords)%nat. - Variable MaxQueries : nat. - Hypothesis MaxQueries_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length q <= MaxQueries)%nat. - - Variable MaxIDs : nat. - Hypothesis MaxIDs_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length db <= MaxIDs)%nat. - - Print XSetSetup_2. - - Theorem XSet_length_2 : - forall k_X k_I db sigmas q s, - In (db, q, s) (getSupport A1) -> - (length (XSetSetup_2 k_X k_I db sigmas) <= MaxKeywords * MaxMatchingIDs)%nat. - intuition. - unfold XSetSetup_2. - Admitted. - - Theorem G10_3_bad_small : - Pr[G10_3_bad] <= (((MaxKeywords + MaxQueries) * MaxMatchingIDs)^2) * - MaxQueries * MaxMatchingIDs * MaxIDs * - MaxKeywords * MaxMatchingIDs / 2 ^ lambda. - - unfold G10_3_bad. - - do 3 - (inline_first; - comp_irr_l; wftac). - - comp_simp. - inline_first. - comp_irr_l; wftac. - eapply compMap_wf. - intuition. - unfold sampleSigmas_11. - comp_simp. - - eapply RndPerm_wf. - - Theorem oc_compMap_qam : - forall (A B C D : Set)(eqdd : EqDec D)(c : A -> OracleComp B C D)(ls : list A) n, - (forall a, In a ls -> queries_at_most (c a) n) -> - queries_at_most (oc_compMap _ c ls) (length ls * n). - - induction ls; intuition; simpl in *. - econstructor. - - econstructor. - eapply H. - intuition. - - intuition. - econstructor. - econstructor. - eapply IHls. - intuition. - intuition. - econstructor. - omega. - - Qed. - - Theorem Initialize_indLoop_11_qam : - forall x0 c' a0 a, - queries_at_most (Initialize_indLoop_11 x0 c' a0 a) 1. - - intuition. - unfold Initialize_indLoop_11. - econstructor. - econstructor. - econstructor. - intuition. - comp_simp. - econstructor. - econstructor. - intuition. - econstructor. - omega. - - Qed. - - Theorem Initialize_wLoop_11_qam : - forall d x0 a, - (exists y z, In (d, y, z) (getSupport A1)) -> - queries_at_most (Initialize_wLoop_11 d x0 a) MaxMatchingIDs. - - intuition. - unfold Initialize_wLoop_11. - comp_simp. - econstructor. - econstructor. - econstructor. - intuition. - econstructor. - eapply oc_compMap_qam. - intuition. - eapply Initialize_indLoop_11_qam. - intuition. - econstructor. - rewrite combine_length. - rewrite allNatsLt_length. - rewrite plus_0_l. - rewrite plus_0_r. - rewrite mult_1_r. - eapply le_trans. - eapply Min.le_min_r. - destruct H0. - destruct H. - eapply MaxMatchingIDs_correct. - eauto. - Qed. - - Theorem GenTrans_10_1_qam : - forall a x a1 a2 a3, - queries_at_most (GenTrans_10_1 a x (a1, (a2, a3))) (length a3). - - intuition. - unfold GenTrans_10_1. - comp_simp. - econstructor. - econstructor. - eapply oc_compMap_qam. - intuition. - econstructor. - econstructor. - intuition. - econstructor. - intuition. - econstructor. - rewrite allNatsLt_length. - omega. - - Qed. - - Theorem Initialize_10_2_qam : - forall x x0 d l x1 xSet, - (exists z, In (d, l, z) (getSupport A1)) -> - queries_at_most (Initialize_10_2 x x0 d l x1 xSet) - ((length (toW d) + length l) * MaxMatchingIDs)%nat. - - intuition. - unfold Initialize_10_2. - econstructor. - econstructor. - eapply oc_compMap_qam. - intuition. - eapply Initialize_wLoop_11_qam. - destruct H0. - econstructor. - econstructor. - eauto. - - intuition. - remember (split c') as z. - destruct z. - comp_simp. - - econstructor. - econstructor. - intuition. - comp_simp. - econstructor. - econstructor. - intuition. - econstructor. - econstructor. - eapply oc_compMap_qam. - intuition. - destruct a. - destruct p0. - econstructor. - eapply GenTrans_10_1_qam. - - repeat (destruct H4). - simpl in H4. - repeat simp_in_support. - eapply (@compMap_support _ _ (fun a b => length (snd b) <= MaxMatchingIDs))%nat in H6. - eapply list_pred_fst_split_irr_if in H6. - assert (l2 = snd (t1, l2)). - trivial. - rewrite H3. - eapply (@list_pred_impl_all _ _ _ (fun x => length (snd x) <= MaxMatchingIDs)%nat). - eapply list_pred_impl. - eapply H6. - intuition. - intuition. - eapply H7. - eapply in_combine_r. - eauto. - - intuition. - repeat simp_in_support. - simpl. - - Show. - repeat (destruct H1). - eapply (@oc_compMap_support _ _ _ (fun s a b => length (fst b) <= MaxMatchingIDs))%nat in H1. - - Theorem arrayLookupList_all : - forall (A B: Set)(eqda : EqDec A)(P : list B -> Prop)(arr : list (A * list B)) a, - P nil -> - (forall ls, In ls (snd (split arr)) -> P ls) -> - P (arrayLookupList _ arr a). - - intuition. - unfold arrayLookupList. - case_eq (arrayLookup eqda arr a); intuition. - eapply H0. - eapply arrayLookup_Some_In. - eauto. - - Qed. - - Show. - eapply list_pred_symm in H1. - SearchAbout list_pred split. - eapply (@list_pred_fst_split_irr _ _ _ (fun x _ => length x <= MaxMatchingIDs)%nat) in H1. - rewrite <- Heqz in H1. - simpl in H1. - eapply list_pred_symm in H1. - eapply arrayLookupList_all. - simpl. - intuition. - intuition. - eapply (@list_pred_impl_all _ _ _ (fun x => length x <= MaxMatchingIDs)%nat). - eapply list_pred_impl. - eapply H1. - intuition. - intuition. - eapply H10. - - Theorem In_snd_split_combine : - forall ( A B : Set)(lsa : list A)(lsb : list B) b, - In b (snd (split (combine lsa lsb))) -> - In b lsb. - - induction lsa; intuition; simpl in *. - intuition. - - destruct lsb; simpl in *. - intuition. - remember (split (combine lsa lsb)) as z. - destruct z. - simpl in *. - intuition. - right. - eapply IHlsa; eauto. - rewrite <- Heqz. - trivial. - Qed. - - eapply In_snd_split_combine. - eauto. - - intuition. - unfold Initialize_wLoop_11 in *. - simpl in H9. - repeat simp_in_support. - destruct x12. - repeat simp_in_support. - simpl. - eapply oc_compMap_length in H11. - rewrite H11. - rewrite combine_length. - rewrite Nat.le_min_r. - rewrite allNatsLt_length. - destruct H0. - eapply MaxMatchingIDs_correct. - eauto. - - intuition. - rewrite combine_length. - - assert (length l = length c'0). - repeat (destruct H4). - simpl in H4. - repeat simp_in_support. - eapply compMap_length in H5. - rewrite H5. - rewrite split_length_l. - trivial. - rewrite <- H5. - rewrite Min.min_idempotent. - intuition. - - intuition. - econstructor. - - repeat rewrite plus_0_l. - rewrite plus_0_r. - rewrite Max.max_0_l. - rewrite mult_plus_distr_r. - eapply plus_le_compat; intuition. - eapply mult_le_compat; intuition. - rewrite combine_length. - eapply Nat.le_min_l. - Qed. - - Theorem Initialize_10_2_badBlinding_prob : - forall x x0 d l x1 xSet, - (exists z, In (d, l, z) (getSupport A1)) -> - Pr[z <-$ (Initialize_10_2 x x0 d l x1 xSet) _ _ natRandFunc nil; - ret ( badBlinding x0 x d (snd z) xSet l)] <= - (((length (toW d) + length l) * MaxMatchingIDs)^2) * - length l * MaxMatchingIDs * length d * - length xSet / 2 ^ lambda. - - intuition. - - eapply leRat_trans. - eapply (@oc_eventProb_0_1 _ (fun f => length f) (fun f => badBlinding x0 x d f xSet l) - (fun n => length l * MaxMatchingIDs * length d * n * length xSet / 2 ^ lambda)). - - eapply Initialize_10_2_qam. - trivial. - - admit. - - Theorem badBlinding_c_nil_f_false : - forall x z xSet a0 a, - badBlinding_c x z nil xSet a0 a = false. - - intuition. - unfold badBlinding_c. - comp_simp. - simpl. - trivial. - - Qed. - - Theorem badBlinding_q_nil_f_false_h : - forall ls x z xSet a, - fold_left - (fun (b : bool) (z0 : nat) => - if b then true else badBlinding_c x z nil xSet a z0) - ls false = false. - - induction ls; intuition; simpl in *. - rewrite badBlinding_c_nil_f_false. - eapply IHls. - - Qed. - - Theorem badBlinding_q_nil_f_false : - forall x z xSet a, - badBlinding_q x z nil xSet a = false. - - intuition. - unfold badBlinding_q. - eapply badBlinding_q_nil_f_false_h. - - Qed. - - Theorem badBlinding_nil_f_false : - forall x0 x d xSet l, - badBlinding x0 x d nil xSet l = false. - - induction l; intuition. - unfold badBlinding, setLet. - - simpl. - rewrite badBlinding_q_nil_f_false. - eapply IHl. - - Qed. - - eapply badBlinding_nil_f_false . - - intuition. - - unfold natRandFunc, randomFunc. - case_eq ( @arrayLookup Blist nat (@list_EqDec bool bool_EqDec) s a); intuition. - comp_simp. - simpl. - rewrite H1. - rewrite evalDist_ret_0; intuition. - eapply rat0_le_all. - comp_inline_l. - comp_at comp_ret leftc 1%nat. - simpl. - - - Fixpoint removeAll (A : Set) (eqd : eq_dec A) (ls : list A) (a : A) := - match ls with - | nil => nil - | a' :: ls' => if eqd a a' then (removeAll eqd ls' a) else a' :: removeAll eqd ls' a - end. - - Theorem removeAll_length_ls : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(a : A), - (length (removeAll eqd ls a) <= length ls)%nat. - - induction ls; intuition; simpl in *. - destruct (eqd a0 a). - subst. - eapply le_S. - eapply IHls; intuition. - simpl. - eapply le_n_S. - eapply IHls; intuition. - Qed. - - Theorem removeAll_In_length : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(a : A), - In a ls -> - (S (length (removeAll eqd ls a)) <= length ls)%nat. - - induction ls; intuition; simpl in *. - intuition. - intuition; subst. - destruct (eqd a0 a0); intuition. - eapply le_n_S. - eapply removeAll_length_ls. - - destruct (eqd a0 a); intuition. - simpl. - eapply le_n_S. - eapply IHls; intuition. - - Qed. - - Theorem badBlinding_ezxz_count : - forall ys xSet e z xind, - NoDup ys -> - (sumList ys (fun y => if (badBlinding_ezxz xSet e z (xind, y)) then 1 else 0) <= (length xSet) / 1). - - induction ys; intuition. - - unfold sumList; simpl. - eapply rat0_le_all. - - rewrite sumList_cons. - inversion H; clear H; subst. - case_eq (badBlinding_ezxz xSet e z (xind, a)); intuition. - unfold badBlinding_ezxz in H. - destruct (eq_nat_dec a z); intuition. - discriminate. - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (e * z * xind * inverse a)) xSet). - remember ((removeAll (EqDec_dec _) xSet (g ^ (e * z * xind * inverse a)))) as xSet'. - assert ((1 + length xSet') <= length xSet)%nat. - simpl. - subst. - eapply removeAll_In_length; intuition. - eapply leRat_trans. - Focus 2. - eapply (leRat_terms (pos 1) (pos 1) H0). - intuition. - clear H0. - - eapply leRat_trans. - Focus 2. - eapply eqRat_impl_leRat. - eapply eqRat_symm. - eapply ratAdd_num. - eapply ratAdd_leRat_compat. - intuition. - - Theorem In_removeAll_ne : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(a b : A), - In a ls -> - a <> b -> - In a (removeAll eqd ls b). - - induction ls; intuition; simpl in *. - intuition; subst. - destruct (eqd b a0). - subst. - intuition. - - simpl; intuition. - - destruct (eqd b a); subst. - eapply IHls; intuition. - simpl. - right. - eapply IHls; intuition. - - Qed. - - Theorem In_removeAll_if : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(a b : A), - In a (removeAll eqd ls b) -> - In a ls. - - induction ls; intuition; simpl in *. - destruct (eqd b a); subst. - right. - eapply IHls; intuition. - eauto. - - simpl in *. - intuition. - right. - eapply IHls; eauto. - - Qed. - - Theorem badBlinding_ezxz_remove_equiv : - forall ys xSet e z xind a, - In (g ^ (e * z * xind * inverse a)) xSet -> - ~ In a ys -> - sumList ys - (fun y : nat => if badBlinding_ezxz xSet e z (xind, y) then 1 else 0) == - sumList ys - (fun y : nat => if badBlinding_ezxz (removeAll (EqDec_dec nat_EqDec) xSet - (g ^ (e * z * xind * inverse a))) e z (xind, y) then 1 else 0). - - - intuition. - eapply sumList_body_eq. - intuition. - - unfold badBlinding_ezxz. - destruct (eq_nat_dec a0 z); intuition. - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (e * z * xind * inverse a0)) xSet). - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (e * z * xind * inverse a0)) - (removeAll (EqDec_dec nat_EqDec) xSet - (g ^ (e * z * xind * inverse a)))). - intuition. - exfalso. - eapply n0. - eapply In_removeAll_ne. - trivial. - - (* need a bit of group theory to prove this*) - admit. - - destruct (in_dec (EqDec_dec nat_EqDec) (g ^ (e * z * xind * inverse a0)) - (removeAll (EqDec_dec nat_EqDec) xSet - (g ^ (e * z * xind * inverse a)))); intuition. - exfalso. - eapply n0. - eapply In_removeAll_if. - eauto. - Qed. - - rewrite badBlinding_ezxz_remove_equiv . - subst. - eapply IHys; intuition. - trivial. - intuition. - - discriminate. - rewrite <- ratAdd_0_l. - eapply IHys; intuition. - Qed. - - Theorem badBlinding_ez_count : - forall ys a xSet e z xinds f, - badBlinding_ez xinds f xSet e z = false -> - NoDup ys -> - (sumList ys (fun y => if (badBlinding_ez xinds ((a, y) :: f) xSet e z) then 1 else 0) <= (length xinds * length xSet) / 1). - - intuition. - unfold badBlinding_ez. - unfold setLet. - - assert (forall y, - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (allPairs_2 xinds (snd (split ((a, y) :: f)))) false - = - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (map (fun z => (z, y)) xinds) false - ). - - Theorem badBlinding_ez_allPairs_simp : - forall xinds f xSet e z, - badBlinding_ez xinds f xSet e z = false -> - forall a y, - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (allPairs_2 xinds (snd (split ((a, y) :: f)))) false = - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (foreach (z0 in xinds)(z0, y)) false. - - intuition. - unfold badBlinding_ez in *. - unfold setLet in *. - simpl. - remember (split f) as w. - destruct w. - simpl in *. - - Theorem fold_left_one_bad : - forall (A : Set)(ls : list A)(f : A -> bool)(a : A), - In a ls -> - f a = true -> - fold_left (fun (b : bool) z => if b then true else (f z)) ls false = true. - - induction ls; intuition; simpl in *. - intuition; subst. - rewrite H0. - eapply fold_left_bad_true. - - case_eq (f a); intuition. - eapply fold_left_bad_true. - eapply IHls; - eauto. - Qed. - - Theorem fold_left_bad_perm_equiv : - forall (A : Set)(ls1 ls2 : list A)(f : A -> bool), - (forall a, f a = true -> (In a ls1 <-> In a ls2)) -> - fold_left (fun (b : bool) a => if b then true else (f a)) ls1 false = - fold_left (fun (b : bool) a => if b then true else (f a)) ls2 false. - - - induction ls1. - induction ls2; intuition; simpl in *. - - case_eq (f a); intuition. - specialize (H a); intuition. - eapply IHls2; intuition. - eapply H. - eapply H1; intuition. - intuition. - - intuition. - simpl in *. - case_eq (f a); intuition. - rewrite fold_left_bad_true. - erewrite fold_left_one_bad. - trivial. - eapply H. - eauto. - intuition. - trivial. - - eapply IHls1. - intuition. - eapply H; intuition. - specialize (H a0); intuition. - subst. - congruence. - - Qed. - - assert (fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (allPairs_2 xinds (y :: l0)) false - = - fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (allPairs_2 xinds l0 ++ (foreach (z0 in xinds)(z0, y)) ) false - ). - - eapply (fold_left_bad_perm_equiv) . - - intuition. - unfold allPairs_2, allPairs_1 in *. - simpl in *. - eapply in_flatten in H1. - destruct H1. - intuition. - eapply in_map_iff in H2. - destruct H2. - intuition; subst. - simpl in *. - intuition. - subst. - eapply in_or_app. - right. - eapply in_map_iff. - econstructor. - intuition. - - eapply in_map_iff in H1. - destruct H1. - intuition. - subst. - - eapply in_or_app. - left. - eapply in_flatten. - econstructor. - split. - eapply in_map_iff. - econstructor. - intuition. - eauto. - eapply in_map_iff. - econstructor. - intuition. - - unfold allPairs_2, allPairs_1 in *. - simpl in *. - eapply in_app_or in H1. - intuition. - - eapply in_flatten in H2. - destruct H2. - intuition. - eapply in_map_iff in H2. - destruct H2. - intuition. - subst. - eapply in_map_iff in H3. - destruct H3. - intuition. - subst. - eapply in_flatten. - econstructor. - split. - eapply in_map_iff. - econstructor. - split. - eauto. - eauto. - simpl. - right. - eapply in_map_iff. - econstructor. - intuition. - eapply in_map_iff in H2. - destruct H2. - intuition. - subst. - eapply in_flatten. - econstructor. - split. - eapply in_map_iff. - econstructor. - split. - eauto. - eauto. - simpl. - intuition. - - rewrite H0. - clear H0. - rewrite fold_left_app. - rewrite H. - trivial. - - Qed. - - eapply badBlinding_ez_allPairs_simp. - trivial. - - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply sumList_body_eq. - intros. - rewrite (H1). - - eapply eqRat_refl. - - - Theorem badBlinding_ez_count_h : - forall xinds ys xSet e z, - NoDup ys -> - sumList ys - (fun a0 : nat => - if fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (foreach (z0 in xinds)(z0, a0)) false - then 1 - else 0) <= (length xinds * length xSet) / 1. - - induction xinds; intuition. - simpl in *. - eapply eqRat_impl_leRat. - eapply sumList_0. - intuition. - - Local Opaque badBlinding_ezxz. - simpl. - - eapply leRat_trans. - eapply sumList_le. - intros. - - assert ( - (if fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (foreach (z0 in xinds)(z0, a0)) (badBlinding_ezxz xSet e z (a, a0)) - then 1 - else 0) - <= - (if (badBlinding_ezxz xSet e z (a, a0)) then 1 else 0) + - (if fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e z x) - (foreach (z0 in xinds)(z0, a0)) false - then 1 - else 0) - )%rat. - - case_eq (badBlinding_ezxz xSet e z (a, a0)); intuition. - rewrite fold_left_bad_true. - - assert (1 <= 1 + 0)%rat. - eapply eqRat_impl_leRat. - eapply ratAdd_0_r. - eapply leRat_trans. - eapply H2. - clear H2. - - eapply ratAdd_leRat_compat; intuition. - eapply rat0_le_all. - - eapply eqRat_impl_leRat. - eapply ratAdd_0_l. - eapply H1. - - rewrite sumList_sum. - eapply leRat_trans. - Focus 2. - eapply eqRat_impl_leRat. - eapply eqRat_symm. - eapply ratAdd_num. - - eapply ratAdd_leRat_compat. - eapply badBlinding_ezxz_count. - trivial. - - eapply IHxinds. - trivial. - - Qed. - - eapply badBlinding_ez_count_h. - trivial. - Qed. - - - Theorem badBlinding_c_count : - forall ys k_X xinds f xSet q c a, - badBlinding_c k_X xinds f xSet q c = false -> - NoDup ys -> - (sumList ys (fun y => if (badBlinding_c k_X xinds ((a, y) :: f) xSet q c) then 1 else 0) <= (length xinds * length f * length xSet) / 1). - - intuition. - unfold badBlinding_c in *. - unfold setLet in *. - destruct q. - - case_eq (eqb (k ++ c) a); intuition. - assert ( - sumList ys - (fun y : nat => - if match - arrayLookup (list_EqDec bool_EqDec) ((a, y) :: f) (k ++ c) - with - | Some z => badBlinding_ez xinds ((a, y) :: f) xSet (F_P k_X k0) z - | None => false - end - then 1 - else 0) - == - sumList ys - (fun y : nat => - if (badBlinding_ez xinds ((a, y) :: f) xSet (F_P k_X k0) y) - - then 1 - else 0) - ). - - eapply sumList_body_eq; intuition. - unfold arrayLookup. - rewrite H. - intuition. - - rewrite H3. - clear H3. - - Theorem badBlinding_ez_count_z : - forall ys xSet e xinds f, - NoDup ys -> - (sumList ys (fun y => if (badBlinding_ez xinds f xSet e y) then 1 else 0) <= (length xinds * length f * length xSet) / 1). - - intuition. - unfold badBlinding_ez. - unfold setLet. - - (* - Theorem sumList_count_orb_le : - forall (A : Set)(ls : list A)(f1 f2 : A -> bool), - sumList ls (fun a => if ((f1 a) || (f2 a)) then 1 else 0) <= - sumList ls (fun a => (if (f1 a) then 1 else 0) + (if (f2 a) then 1 else 0))%rat. - - induction ls; intuition. - unfold sumList; simpl. - intuition. - - repeat rewrite sumList_cons. - eapply leRat_trans. - Focus 2. - eapply eqRat_impl_leRat. - eapply eqRat_symm. - eapply sumList_cons. - simpl. - destruct (f1 a); simpl. - eapply ratAdd_leRat_compat; intuition. - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply ratAdd_0_r. - eapply ratAdd_leRat_compat; intuition. - eapply rat0_le_all. - - destruct (f2 a). - eapply ratAdd_leRat_compat; intuition. - eapply eqRat_impl_leRat. - eapply ratAdd_0_l. - - eapply ratAdd_leRat_compat; intuition. - eapply rat0_le_all. - - Qed. - *) - - Theorem badBlinding_ezxz_count_z : - forall ys xSet e p, - NoDup ys -> - (sumList ys (fun y => if (badBlinding_ezxz xSet e y p) then 1 else 0) <= (length xSet) / 1). - - Admitted. - - Theorem badBlinding_ez_count_z_h : - forall ps ys xSet e, - NoDup ys -> - sumList ys - (fun y : nat => - if fold_left - (fun (b : bool) (x : nat * nat) => - if b then true else badBlinding_ezxz xSet e y x) - ps false - then 1 - else 0) <= length ps * length xSet / 1. - - induction ps; intuition; simpl in *. - eapply eqRat_impl_leRat. - eapply sumList_0. - intuition. - - assert ( sumList ys - (fun y : nat => - if fold_left - (fun (b0 : bool) (x : nat * nat) => - if b0 then true else badBlinding_ezxz xSet e y x) ps - (badBlinding_ezxz xSet e y (a0, b)) - then 1 - else 0) - <= - sumList ys - (fun y : nat => - (if (badBlinding_ezxz xSet e y (a0, b)) then 1 else 0)%rat + - if fold_left - (fun (b0 : bool) (x : nat * nat) => - if b0 then true else badBlinding_ezxz xSet e y x) ps - false - then 1 - else 0)%rat - ). - - eapply sumList_le; intuition. - destruct (badBlinding_ezxz xSet e a (a0, b)). - rewrite fold_left_bad_true. - - assert (1 == 1 + 0)%rat. - eapply ratAdd_0_r. - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply H1. - eapply ratAdd_leRat_compat; intuition. - eapply rat0_le_all. - - eapply eqRat_impl_leRat. - eapply ratAdd_0_l. - rewrite H0. - clear H0. - rewrite sumList_sum. - eapply leRat_trans. - Focus 2. - eapply eqRat_impl_leRat. - symmetry. - eapply ratAdd_num. - eapply ratAdd_leRat_compat; intuition. - - eapply badBlinding_ezxz_count_z. - trivial. - Qed. - - eapply leRat_trans. - eapply badBlinding_ez_count_z_h. - trivial. - - Theorem allPairs_2_length : - forall (A B : Set)(lsa : list A)(lsb : list B), - (length (allPairs_2 lsa lsb) = length lsa * length lsb)%nat. - - induction lsa; intuition. - unfold allPairs_2, allPairs_1 in *; simpl. - rewrite app_length. - - rewrite IHlsa. - rewrite map_length. - intuition. - Qed. - - rewrite allPairs_2_length. - rewrite split_length_r. - intuition. - - Qed. - - (* In order to apply the next theorem, we need to show that putting (a, y) in the function changes nothing (since y is also used as the unblinding values) *) - assert - (sumList ys - (fun y : nat => - if badBlinding_ez xinds ((a, y) :: f) xSet (F_P k_X k0) y then 1 else 0) - == - sumList ys - (fun y : nat => - if badBlinding_ez xinds f xSet (F_P k_X k0) y then 1 else 0) - ). - admit. - rewrite H3. - clear H3. - - eapply badBlinding_ez_count_z; intuition. - - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply sumList_body_eq. - intuition. - simpl. - - case_eq (@eqb Blist (@list_EqDec bool bool_EqDec) - (@app bool k (natToBlist c)) a); intuition. - rewrite eqb_leibniz in H4. - subst. - rewrite eqb_refl in H. - discriminate. - eapply eqRat_refl. - - case_eq (@arrayLookup Blist nat (@list_EqDec bool bool_EqDec) f - (@app bool k (natToBlist c))); intuition. - rewrite H3 in H0. - eapply leRat_trans. - eapply badBlinding_ez_count; intuition. - eapply leRat_terms. - eapply mult_le_compat; intuition. - rewrite <- mult_1_r at 1. - eapply mult_le_compat; intuition. - (* the size of the function is at least 1 *) - admit. - intuition. - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply sumList_0. - intuition. - eapply rat0_le_all. - Qed. - - Theorem badBlinding_q_count : - forall ys k_X xinds f xSet q a, - badBlinding_q k_X xinds f xSet q = false -> - NoDup ys -> - (sumList ys (fun y => if (badBlinding_q k_X xinds ((a, y) :: f) xSet q) then 1 else 0) <= (MaxMatchingIDs * length xinds * length f * length xSet) / 1). - Admitted. - - Theorem badBlinding_count : - forall ys k_I k_X db f xSet q a, - badBlinding k_I k_X db f xSet q = false -> - NoDup ys -> - (sumList ys (fun y => if (badBlinding k_I k_X db ((a, y) :: f) xSet q) then 1 else 0) <= (length q * MaxMatchingIDs * length db * length f * length xSet) / 1). - Admitted. - - Theorem badBlinding_prob : - forall k_I k_X db f xSet q a, - badBlinding k_I k_X db f xSet q = false -> - Pr [a0 <-$ [ 0 .. 2 ^ lambda); ret badBlinding k_I k_X db ((a, a0) :: f) xSet q ] <= (length q * MaxMatchingIDs * length db * length f * length xSet) / 2 ^ lambda. - - intuition. - rewrite evalDist_seq_step. - - eapply leRat_trans. - eapply sumList_le. - intuition. - eapply ratMult_leRat_compat. - eapply eqRat_impl_leRat. - eapply RndNat_prob. - eapply RndNat_support_lt. - trivial. - eapply leRat_refl. - rewrite sumList_factor_constant_l. - - assert ( - (@sumList nat - (@getSupport nat (RndNat (pow (S (S O)) (posnatToNat lambda)))) - (fun a0 : nat => - @evalDist bool - (@Ret bool (@EqDec_dec bool bool_EqDec) - (badBlinding k_I k_X db - (@cons (prod Blist nat) (@pair Blist nat a a0) f) xSet q)) - true)) - == - (sumList (allNatsLt (2^ lambda))) - (fun a0 : nat => if (badBlinding k_I k_X db ((a, a0) :: f) xSet q) then 1 else 0) - - ). - - erewrite sumList_permutation. - eapply sumList_body_eq. - intuition. - Local Transparent evalDist. - simpl. - destruct ( badBlinding k_I k_X db ((a, a0) :: f) xSet q). - destruct (EqDec_dec bool_EqDec true true); intuition. - destruct (EqDec_dec bool_EqDec false true); intuition. - - Theorem RndNat_support_perm : - forall (n : nat), - Permutation (getSupport (RndNat n)) (allNatsLt n). - - intuition. - eapply NoDup_Permutation. - eapply getSupport_NoDup. - eapply allNatsLt_NoDup. - intuition. - eapply allNatsLt_lt_if. - eapply RndNat_support_lt. - trivial. - eapply in_getSupport_RndNat. - eapply allNatsLt_lt. - trivial. - - Qed. - - eapply RndNat_support_perm . - - rewrite H. - clear H. - - rewrite badBlinding_count. - eapply eqRat_impl_leRat. - rewrite <- ratMult_num_den. - eapply eqRat_terms. - omega. - unfold posnatMult, natToPosnat, posnatToNat. - eapply mult_1_r. - trivial. - eapply allNatsLt_NoDup. - Qed. - - Show. - rewrite badBlinding_prob. - admit. - - trivial. - admit. - trivial. - - rewrite <- ratMult_num_den. - eapply leRat_terms. - simpl. - eapply le_eq. - ring. - unfold posnatMult, natToPosnat, posnatToNat. - rewrite mult_1_l. - trivial. - - Qed. - - Show. - rewrite Initialize_10_2_badBlinding_prob. - eapply leRat_terms. - (* need to clean up the expression a bit*) - admit. - - intuition. - - econstructor. - eauto. - Qed. - - - Print G11. - Print Initialize_11. - Print GenTrans_11. - Print ServerSearch_11. - Print ServerSear - - - - Theorem G10_3_equiv : - | Pr[x <-$ G10_2] - Pr[G10_3] | <= 0. - - rewrite Initialize_10_2_bad_Prob. - - - rewrite mult_assoc; - eapply mult_le_compat; intuition. - ea - ring. - omega. - - (* TODO: replace MaxMatchingIDs everywhere with (length db)---simplifies things and gets rind of an assumptions, but makes bounds worse.*) - - Print lookupInds. - - SearchAbout ratMult RatIntro. - - - Local Transparent evalDist. - simpl. - unfold evalDist. - - assert ( - sumList (getSupport ([ 0 .. 2 ^ lambda))) - (fun a0 : nat => - Pr [ret badBlinding k_I k_X db ((a, a0) :: f) xSet q ]))%rat - == - sumList (getSupport ([ 0 .. 2 ^ lambda))) - (fun a0 : nat => - Pr [ret badBlinding k_I k_X db ((a, a0) :: f) xSet q ]))%rat - ). - - (sumList ys (fun y => if (badBlinding k_I k_X db ((a, y) :: f) xSet q) then 1 else 0)) - == - (sumList ys (fun y => if (badBlinding k_I k_X db ((a, y) :: f) xSet q) then 1 else 0)) - ). - - - eapply leRat_trans. - eapply ratMult_leRat_compat. - eapply leRat_refl. - Focus 2. - rewrite badBlinding_count. - - - Qed. - - Set Printing All. - Show. - case_eq (arrayLookup (list_EqDec bool_EqDec) s a ); intuition. - - - - SearchAbout badBlinding nil - - - unfold Initialize_10_2. - - - Qed. - - Show. - eapply Initialize_10_2_badBlinding_prob. - - Show. - comp_inline_l. - inline_first. - - - Qed. - - -Here!!! - - - intuition. - - unfold Initialize_10_2, Initialize_11. - simpl. - - assert TSet. - assert (TSet * TSetKey). - eapply comp_base_exists. - eapply TSetSetup. - econstructor. - intuition. - - comp_skip. - - remember (split a0) as z. - destruct z. - - simpl. - inline_first. - comp_skip. - - simpl. - comp_simp. - simpl. - inline_first. - comp_skip. - comp_simp. - - case_eq (badBlinding k_I k_X db b xSet q); intuition. - comp_irr_l. - admit. - admit. - comp_irr_r. - admit. - admit. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - - - - erewrite natRandFunc_badBlinding_persists_oc. - symmetry. - erewrite natRandFunc_badBlinding_persists_oc. - trivial. - eauto. - eapply H13. - eauto. - eapply H12. - simpl in *. - assert (badBlinding k_I k_X db l2 xSet q = true). - eapply natRandFunc_badBlinding_persists_oc. - eauto. - eauto. - congruence. - - comp_skip. - (* This can go bad during the course of execution. *) - - - - - - - Show. - - admit. - - intuition. - subst. - eapply comp_spec_consequence. - eapply GenTrans_10_1_11_spec. - apply in_combine_r in H15. - eapply (@compMap_support _ _ (fun a b => snd b = arrayLookupList _ (combine (toW db) l) a)) in H9. - eapply list_pred_symm in H9. - eapply (@list_pred_impl_unary _ _ _ (fun a a => (snd a) = nil \/ In (snd a) l)) in H9. - - - - eapply list_pred_unary_in in H9. - intuition. - assert (snd (a2, b3) = nil). - eapply H14. - simpl in *. - subst. - simpl. - omega. - simpl in *. - - eapply (@oc_compMap_support _ _ _ (fun s a b => length (snd a) <= MaxMatchingIDs -> length (fst b) <= MaxMatchingIDs))%nat in H5. - eapply list_pred_symm in H5. - - eapply (@list_pred_fst_split_irr _ _ _ (fun - (b0 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) - (a : Keyword * list nat) => - (length (snd a) <= MaxMatchingIDs -> length b0 <= MaxMatchingIDs)%nat)) in H5. - - - - eapply list_pred_symm in H5. - eapply (@list_pred_snd_split_irr _ _ _ (fun - (a : list nat) - (b0 : list - (Vector.t bool (lambda + lambda) * - (Identifier lambda * nat))) - => - (length a <= MaxMatchingIDs -> length b0 <= MaxMatchingIDs)%nat)) in H5. - - eapply (@list_pred_in_all _ _) in H5. - eapply H5. - intuition. - eapply H1. - assert (split (combine (toW db) sigmas) = ((toW db), sigmas)). - eapply combine_split. - trivial. - rewrite H17 in H9. - trivial. - rewrite <- Heqz. - trivial. - - intuition. - simpl in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - simpl. - - - - erewrite (@oc_compMap_length _ _ _ _ _ _ _ (list (Blist * nat))). - Focus 2. - eapply H19. - rewrite combine_length. - eapply Nat.min_le_iff. - right. - rewrite allNatsLt_length. - eauto. - - intuition. - trivial. - intuition. - simpl in *. - unfold arrayLookupList in *. - case_eq - ( @arrayLookup Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (@list_EqDec bool bool_EqDec) - (@combine Keyword - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (prod (Identifier (posnatToNat lambda)) nat))) - (toW db) l) b4); intuition. - rewrite H in H14. - subst. - - - - right. - eapply (@arrayLookup_combine_Some_In Blist). - eauto. - rewrite H in H14. - subst. - intuition. - - intuition. - repeat simp_in_support. - trivial. - - intuition. - assert (c = (snd (snd (a3, (b2, c))))). - trivial. - rewrite H17. - clear H17. - eapply H13. - eapply in_combine_r. - eauto. - trivial. - - assert (In b2 (fst (split db))). - admit. - eauto. - - assert (bvectorToNat (snd (splitVector lambda lambda a3)) = - F_P k_I (Vector.to_list b2) * inverse c). - (* I can derive this from the existing support information *) - admit. - eapply H17. - - eauto. - eapply in_combine_l. - eauto. - - intuition. - simpl. - assert ((a3, (a4, b5)) = fst b4). - destruct b4; simpl in *; pairInv; trivial. - eapply H18. - destruct b4; simpl in *; pairInv; trivial. - simpl. - subst. - - eapply H13. - - eapply H19. - - - Theorem badBlinding_all_false : - forall q k_I k_X db s0 xSet, - badBlinding k_I k_X db s0 xSet q = false -> - forall a, In a q -> - badBlinding_q k_X (map (fun x => F_P k_I (Vector.to_list x)) (fst (split db))) s0 xSet a = false. - - induction q; unfold badBlinding; intuition; simpl in *; - intuition. - subst. - unfold setLet in *. - case_eq ((badBlinding_q k_X - (foreach (x in fst (split db))F_P k_I (Vector.to_list x)) s0 - xSet a0)) - ; intuition. - rewrite H1 in H0. - rewrite fold_left_bad_true in H0. - discriminate. - - unfold setLet in *. - case_eq (badBlinding_q k_X - (foreach (x in fst (split db))F_P k_I (Vector.to_list x)) s0 - xSet a); intuition; - rewrite H1 in H0. - rewrite fold_left_bad_true in H0. - discriminate. - eapply IHq; eauto. - Qed. - - Show. - - Theorem badBlinding_q_xind_subset : - forall x1 x2 k_X s xSet q, - badBlinding_q k_X x1 s xSet q = false -> - (forall a, In a x1 -> In a x2) -> - badBlinding_q k_X x2 s xSet q = false. - - Admitted. - - eapply badBlinding_q_xind_subset. - eapply badBlinding_all_false . - eapply H4. - eapply in_combine_l. - eauto. - - intuition. - (* Combine this with the support fact from before *) - admit. - - intuition. - simpl in *. - - intros. - simpl in *. - destruct H14. - intuition. - simpl in *. - - - SearchAbout badBlinding_q. - eapply goodBlinding_Prop_correct in H4. - unfold badBlinding in *. - unfold setLet in *. - - - SearchAbout arrayLookup Some In. - Set Printing All. - Show. - case_eq (arrayLookup (list_EqDec bool_EqDec) (combine (toW db) l) b4); intuition. - unfold Blist in *. - rewrite H in H4. - - eapply MaxMatchingIDs_correct. - - SearchAbout min le. - SearchAbout combine length. - - eapply (@oc_compMap_support _ _ _ (fun s a b => length b <= MaxMatchingIDs))%nat. - - assert (length (fst (split a0)) <= MaxMatchingIDs)%nat. - eapply H3. - - assert (forall x, In x l -> length x <= MaxMatchingIDs)%nat. - admit. - apply in_combine_r in H10. - - - assert (forall x, In x b1 -> length (snd x) <= MaxMatchingIDs)%nat. - intuition. - - - assert (b3 = snd (a2, b3)). - trivial. - rewrite H13. - clear H13. - eapply H12. - eapply in_combine_r. - eauto. - - - eapply H12. - eapply in_combine_r in H10. - - - Show. - - - simpl. - - eapply H11. - - rewrite H11 in H9. - - - Set Printing All. - Show. - case_eq (arrayLookup (list_EqDec bool_EqDec) s (k ++ a) ); intuition. - rewrite H9 in H10. - - - SearchAbout groupExp mult. - - erewrite <- H2. - - intros. - - - eapply H2. - - - SearchAbout In split. - - eapply H. - - SearchAbout In allNatsLt. - unfold goodBlinding_q_Prop in *. - - - - Theorem badBlinding_q_Prop : - forall - k_X (xinds : list nat)(f : list (Blist * nat)) xSet (q : Query), - badBlinding_q k_X xinds f xSet q = false -> - forall z, In z (allNatsLt MaxMatchingIDs) -> - badBlinding_c k_X xinds f xSet q z = false. - - - fold_left (fun (b : bool) z => if b then true else badBlinding_c k_X xinds f xSet q z) (allNatsLt MaxMatchingIDs) false. - - - - unfold badBlinding_q in H6. - - - subst. - - - comp_simp. - - (forall n0 n1, - - SearchAbout map eq. - eapply map_ext'. - - f_equal. - f_equal. - - - - Qed. - - Qed. - - - eapply Init_10_2_11_eq_until_bad. - comp_simp. - intuition; simpl in *. - destruct( badBlinding x1 x0 d b (XSetSetup_2 x0 x1 d x2) l); intuition. - rewrite <- H8. - comp_irr_l; wftac. - comp_irr_r; wftac. - eapply comp_spec_ret; intuition. - pairInv. - pairInv. - pairInv. - comp_skip. - eapply comp_spec_ret; intuition. - pairInv. - rewrite <- H8. - trivial. - pairInv. - trivial. - - - Qed. - - Theorem G10_4_equiv : - | Pr[x <-$ G10_3; ret fst x] - Pr[x <-$ G10_4; ret fst x] | <= 0. - - eapply leRat_trans. - eapply fundamental_lemma_h. - - - - Qed. - - - (* Now remove the badness computation to get to game 11 *) - - - (* Step 10.2 : Test for badness at the end of the game the result *) - - (* Step 10.3 : Like 11 with badness exposed. Close to 10.2 by identical until bad *) - - (* These are the same unless there is a collision during unblinding. The random function ensures that the probability of this event is small. *) - (* Use identical until bad with oc_eventProb and other similar theorems. *) - - - - (* Step 12: Replace the random function calls in GenTrans with array lookups. This is equivalent because unblidinding collisions are now impossible. *) - - - (* Step 13: replace random function calls in initialization with random values, since they are all unique*) - - - - - - (* Step 11 : We want to replace the random function with a random permutation that ensures that there are no unbliding collisions in the XSet. First we move some things around to get the game in the form of the "generalized random permutation" argument. *) - Require Import RandPermSwitching. - - Definition Initialize_wLoop_11 db k_I e := - [w, sigma ] <-2 e; - inds <- lookupInds db w; - k_E <--$$ {0, 1}^lambda; - oc_compMap _ (Initialize_indLoop_9 k_I k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))). - - Definition G11_A db q s_A k_X k_I sigmas xSet := - t_entries_pts <--$ oc_compMap _ (Initialize_wLoop_11 db k_I) (combine (toW db) sigmas); - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_9 edb k_X) (combine q stags_ts); - $ A2 s_A edb (snd (split searchRes)). - - - Definition sampleSigmas_11 db w := - inds <- lookupInds db w; - RndPerm (length inds). - - Definition G11 := - [db, q, s_A] <-$3 A1; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [b, _] <-$2 (G11_A db q s_A k_X k_I sigmas xSet) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - ret b. - - - (* Step 12 : now apply the argument *) - - (* first we need the predicate that tells us when the PRP goes bad *) - - Definition allPairs_1(A B : Type)(a : A)(lsb : list B) := - map (pair a) lsb. - - Definition allPairs_2(A B : Type)(lsa : list A)(lsb : list B) : list (A * B) := - flatten (map (fun a => allPairs_1 a lsb) lsa). - - Definition rpBad1 (xSet : list nat)(r1 : nat)(p : nat * nat * nat) := - [xind, e, r2] <-3 p; - if (in_dec (EqDec_dec _) (g ^ r1 * xind * e * r2) xSet) - then true else false. - - Definition rpBad k_X (xSet : list nat)(xinds : list nat)(f : list (Blist * nat))(d : Blist)(r : nat) := - keywords <- d :: (fst (split f)); - (* these aren't keywords, they are keywords appended with numbers, and they don't include query keywords *) - es <- map (F_P k_X) keywords; - xinds_es <- allPairs_2 xinds es; - rs <- snd (split f); - xinds_es_rs <- allPairs_2 xinds_es rs; - bad1 <- fold_left (fun (b : bool) x => if b then true else (rpBad1 xSet r x)) xinds_es_rs false; - rsInverse <- map inverse rs; - xinds_es_rsInverse <- allPairs_2 xinds_es rsInverse; - bad2 <- fold_left (fun (b : bool) x => if b then true else (rpBad1 xSet (inverse r) x)) xinds_es_rsInverse false; - bad1 || bad2. - - Variable maxDB_Keywords : nat. - Variable maxQueries : nat. - Variable maxIDsPerKeyword : nat. - Variable maxDB_Size : nat. - - Hypothesis maxDB_Keywords_correct : - forall db qs s_A, - In (db, qs, s_A) (getSupport A1) -> - (length (toW db) <= maxDB_Keywords)%nat. - - Hypothesis maxIDsPerKeyword_correct : - forall db qs s_A, - In (db, qs, s_A) (getSupport A1) -> - forall w, - (length (lookupInds db w) <= maxIDsPerKeyword)%nat. - - Theorem oc_compMap_qam : - forall (A B C D : Set){eqdd : EqDec D}(ls : list A)(c : A -> OracleComp B C D) n, - (forall a, In a ls -> queries_at_most (c a) n) -> - queries_at_most (oc_compMap _ c ls) (length ls * n). - - induction ls; intuition; simpl in *. - econstructor. - econstructor; intuition. - rewrite <- plus_0_r. - econstructor. - eapply IHls. - intuition. - intuition. - econstructor. - Qed. - - Theorem G11_A_qam : - forall a0 b b0 b1 b2 b3 x, - In (a0, b, x) (getSupport A1) -> - In b3 (getSupport (compMap _ (sampleSigmas_11 a0) (toW a0))) -> - queries_at_most (G11_A a0 b b0 b1 b2 b3 (XSetSetup_2 b1 b2 a0 b3)) (maxDB_Keywords * maxIDsPerKeyword + maxQueries * 2). - - intuition. - unfold G11_A. - econstructor. - econstructor. - eapply oc_compMap_qam. - intuition. - unfold Initialize_wLoop_11. - comp_simp. - econstructor. - econstructor. - intuition. - econstructor. - eapply oc_compMap_qam . - intuition. - unfold Initialize_indLoop_9. - comp_simp. - econstructor. - econstructor. - intuition. - econstructor. - econstructor. - intuition. - econstructor. - rewrite plus_0_l. - rewrite plus_0_r. - rewrite mult_1_r. - - rewrite combine_length. - rewrite allNatsLt_length. - unfold sampleSigmas_11 in *. - apply (@compMap_support _ _ (fun a b => length b = length (lookupInds a0 a))) in H1. - - rewrite permute_length_eq. - rewrite min_r. - eapply maxIDsPerKeyword_correct. - eauto. - eapply le_refl_gen. - - - Theorem list_pred_in_combine : - forall (A B : Set) P (lsa : list A)(lsb : list B), - list_pred P lsa lsb -> - forall a b, - In (a, b) (combine lsa lsb) -> - P a b. - - induction 1; intuition; simpl in *. - intuition. - intuition. - pairInv. - trivial. - - Qed. - - eapply list_pred_in_combine in H1. - symmetry. - eapply H1. - trivial. - intuition. - - unfold setLet in *. - eapply RndPerm_In_support_length. - trivial. - - rewrite plus_0_l. - eapply mult_le_compat; intuition. - rewrite combine_length. - apply compMap_length in H1. - rewrite H1. - rewrite min_l. - eapply maxDB_Keywords_correct. - eauto. - trivial. - - intuition. - comp_simp. - simpl. - rewrite <- plus_0_l. - econstructor. - econstructor. - intuition. - rewrite <- plus_0_l. - econstructor. - econstructor. - intuition. - rewrite <- plus_0_r. - econstructor. - - econstructor. - eapply oc_compMap_qam. - intuition. - unfold GenTrans_9. - comp_simp. - - destruct a1. - destruct p0. - destruct q. - unfold setLet. - - comp_simp. - econstructor. - econstructor. - eapply oc_compMap_qam. - intuition. - econstructor. - econstructor. - intuition. - econstructor. - - - Qed. - - Definition G12 := - [db, q, s_A] <-$3 A1; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - allXinds <- map (F_P k_I) (toW db); - [b, _] <-$2 (G11_A db q s_A k_X k_I sigmas xSet) _ _ (@GenRP (Blist) _ (RndNat (2^lambda)) _ _ (rpBad k_X xSet allXinds)) nil; - ret b. - - Theorem G11_G12_close : - | Pr[G11] - Pr[G12] | <= 0. - - unfold G11, G12. - eapply leRat_trans. - - - eapply evalDist_bind_distance; intuition; wftac. - apply evalDist_bind_distance; intuition; wftac. - apply evalDist_bind_distance; intuition; wftac. - - apply evalDist_bind_distance; intuition; wftac. - admit. - comp_simp. - - destruct a3. - eapply (@RPS_G0_1_close _ _ _ _ _ _ _ _ - (fun (n : nat) => (maxDB_Keywords * (maxDB_Keywords + maxQueries * 2) * maxDB_Size * n) / 2^lambda)). - admit. - - - eapply well_formed_RndNat. - admit. - - SearchAbout RndNat. - - apply evalDist_bind_distance; intuition; wftac. - admit. - - - Qed. - - - Print G9_A. - - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_9 db k_I) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_9 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - - Print G10. - - - - (* Step 7.1: This proof is by a hybrid argument on the list of keywords. Start by delaying some of the sampling *) - Definition Initialize_7_1 (db : DB)(q : list Query) := - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_7 db k_I k_Z) (toW db); - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Theorem G7_1_equiv : - Pr[G7] == Pr[G_gen Initialize_7_1]. - - unfold G7, G_gen, Initialize_7, Initialize_7_1. - - (comp_skip; comp_simp; inline_first). - do 3 (inline_first; comp_at comp_inline leftc 1%nat; comp_swap_l; comp_skip; comp_simp). - repeat (inline_first; comp_skip; comp_simp). - - Qed. - - (* Step 7.2: Now put the game in the form of the hybrid argument. *) - - Require Import ListHybrid. - - Definition G7_2_A1 := - [db, q, s_A] <-$3 A1; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - ret (toW db, (k_I, k_Z, db, q, s_A)). - - Definition G7_2_A2 (s_A : Bvector lambda * Bvector lambda * DB * list Query * A_State) := - let '(k_I, k_Z, db, _, _) := s_A in - Initialize_wLoop_7 db k_I k_Z. - - Definition G7_2_A3 (s : Bvector lambda * Bvector lambda * DB * list Query * A_State) wLoopRes := - let '(k_I, k_Z, db, q, s_A) := s in - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - Definition G7_2 := - [lsa, s] <-$2 G7_2_A1; - lsb <-$ compMap _ (G7_2_A2 s) lsa; - G7_2_A3 s lsb. - - Theorem G7_2_equiv: - Pr[G_gen Initialize_7_1] == Pr[G7_2]. - - unfold G7_2, G_gen, G7_2_A1, G7_2_A2, G7_2_A3, Initialize_7_1. - do 3 (inline_first; comp_skip; comp_simp). - inline_first; comp_skip. - inline_first; comp_skip. - - - remember ( @split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1) as z. - destruct z. - remember (split x1) as z. - destruct z. - - assert (l2 = l0). - assert (l2 = fst (split x1)). - rewrite <- Heqz0. - trivial. - assert (l0 = fst ( ( @split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1))). - rewrite <- Heqz. - trivial. - subst. - trivial. - subst. - - inline_first. - comp_skip. - reflexivity. - - comp_simp. - inline_first. - comp_skip. - reflexivity. - - assert (l1 = l3). - assert (l1 = snd (( @split - (list - (prod - (Bvector - (posnatToNat - (@natToPosnat - (plus (posnatToNat lambda) (posnatToNat lambda)) - (nz_posnat_plus lambda lambda)))) - (Bvector (posnatToNat lambda)))) - (list nat) x1) )). - rewrite <- Heqz. - trivial. - - assert (l3 = snd (split x1)). - rewrite <- Heqz0. - trivial. - subst. - trivial. - subst. - reflexivity. - - Qed. - - Definition G7_3_A2 (s_A : Bvector lambda * Bvector lambda * DB * list Query * A_State) := - let '(k_I, k_Z, db, _, _) := s_A in - Initialize_wLoop_8 db k_I k_Z. - - Definition G7_3 := - [lsa, s] <-$2 G7_2_A1; - lsb <-$ compMap _ (G7_3_A2 s) lsa; - G7_2_A3 s lsb. - - Variable max_keywords : nat. - Hypothesis max_keywords_correct : - forall db q s_A, - In (db, q, s_A) (getSupport A1) -> - (length (toW db) <= max_keywords)%nat. - - Theorem max_keywords_G7_2_A1 : - forall a b, - In (a, b) (getSupport G7_2_A1) -> - (length a <= max_keywords)%nat. - - intuition. - unfold G7_2_A1 in *. - repeat simp_in_support. - destruct x. - destruct p0. - repeat simp_in_support. - eapply max_keywords_correct; eauto. - Qed. - - Require Import Encryption. - - Theorem G7_3_close : - | Pr[G7_2] - Pr[G7_3] | <= 0. - - unfold G7_2, G7_3. - eapply evalDist_bind_distance. - unfold G7_2_A1. - wftac. - intuition. - intuition. - - eapply leRat_trans. - eapply list_hybrid_close. - intuition. - - Theorem G7_2_A2_wf : - forall x y, - well_formed_comp (G7_2_A2 x y). - - Admitted. - - Theorem G7_3_A2_wf : - forall x y, - well_formed_comp (G7_3_A2 x y). - - Admitted. - - eapply G7_3_A2_wf. - - intuition. - unfold LH_G_0, LH_G_1. - comp_simp. - eapply evalDist_bind_distance; intuition. - eapply compMap_wf; intuition. - eapply G7_3_A2_wf. - - eapply leRat_trans. - eapply eqRat_impl_leRat. - eapply ratDistance_eqRat_compat. - comp_swap_l. - reflexivity. - comp_swap_l. - reflexivity. - eapply evalDist_bind_distance; intuition. - eapply compMap_wf; intuition. - eapply G7_2_A2_wf. - - unfold G7_2_A2, G7_3_A2. - - Definition Initialize_indLoop_7_O k_I k_Z w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$ query ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - $ ret (@Vector.append bool lambda lambda e (natToBvector y), ind). - - Definition Initialize_wLoop_7_O db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_7_O k_I k_Z w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Theorem hybrid_enc_equiv0 : - forall n a x z b1 a1 b2 b3 b4, - evalDist - (a4 <-$ Initialize_wLoop_7 b1 a1 b2 (nth n a nil); - G7_2_A3 z (b3 ++ a4 :: b4)) x - == - evalDist - (k <-$ {0, 1}^lambda; - [b, _] <-$2 (a4 <--$ Initialize_wLoop_7_O b1 a1 b2 (nth n a nil); $ G7_2_A3 z (b3 ++ a4 :: b4)) _ _ (EncryptOracle Enc _ k) tt; - ret b) x. - - intuition. - unfold Initialize_wLoop_7. - comp_simp. - comp_inline_l. - comp_at comp_inline leftc 1%nat. - comp_swap_l. - comp_skip. - unfold Initialize_wLoop_7_O. - eapply comp_spec_eq_impl_eq. - simpl. - inline_first; comp_skip. - inline_first. comp_simp. - inline_first. - comp_skip. - - eapply (@compMap_oc_spec _ _ eq _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - - unfold Initialize_indLoop_4, Initialize_indLoop_7_O. - simpl. - unfold EncryptOracle. - inline_first. - comp_skip; try apply (oneVector (lambda + lambda)). - comp_simp. - eapply comp_spec_ret; intuition. - destruct b8. - inline_first. - comp_ret_r. - comp_ret_r. - unfold setLet. - inline_first. - comp_skip. - simpl in H5. - assert (a2 = l). - eapply list_pred_eq_impl_eq. - trivial. - subst. - - remember (@split - (list - (prod (Bvector (plus (posnatToNat lambda) (posnatToNat lambda))) - (Bvector (posnatToNat lambda)))) (list nat) - (@app - (prod - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat)) b5 - (@cons - (prod - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat)) - (@pair - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat) l b7) b6))) as z. - destruct z. - inline_first. - comp_skip. - inline_first. - comp_skip. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - Theorem hybrid_enc_equiv1 : - forall n a x z b1 a1 b2 b3 b4, - evalDist - (a4 <-$ Initialize_wLoop_8 b1 a1 b2 (nth n a nil); - G7_2_A3 z (b3 ++ a4 :: b4)) x - == - evalDist - (k <-$ {0, 1}^lambda; - [b, _] <-$2 (a4 <--$ Initialize_wLoop_7_O b1 a1 b2 (nth n a nil); $ G7_2_A3 z (b3 ++ a4 :: b4)) _ _ (EncryptNothingOracle Enc _ (zeroVector lambda) k) tt; - ret b) x. - - intuition. - unfold Initialize_wLoop_8. - comp_simp. - comp_inline_l. - comp_at comp_inline leftc 1%nat. - comp_swap_l. - comp_skip. - unfold Initialize_wLoop_7_O. - eapply comp_spec_eq_impl_eq. - simpl. - inline_first; comp_skip. - inline_first. comp_simp. - inline_first. - comp_skip. - - eapply (@compMap_oc_spec _ _ eq _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - - unfold Initialize_indLoop_8, Initialize_indLoop_7_O. - simpl. - unfold EncryptNothingOracle. - inline_first. - comp_skip; try apply (oneVector (lambda + lambda)). - comp_simp. - eapply comp_spec_ret; intuition. - destruct b8. - inline_first. - comp_ret_r. - comp_ret_r. - unfold setLet. - inline_first. - comp_skip. - simpl in H5. - assert (a2 = l). - eapply list_pred_eq_impl_eq. - trivial. - subst. - - remember (@split - (list - (prod (Bvector (plus (posnatToNat lambda) (posnatToNat lambda))) - (Bvector (posnatToNat lambda)))) (list nat) - (@app - (prod - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat)) b5 - (@cons - (prod - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat)) - (@pair - (list - (prod - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda))) - (Identifier (posnatToNat lambda)))) - (list nat) l b7) b6))) as z. - destruct z. - inline_first. - comp_skip. - inline_first. - comp_skip. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - rewrite hybrid_enc_equiv0 . - rewrite hybrid_enc_equiv1 . - - inline_first. - - - Set Printing All. - Show. - remember (split (b5 ++ (l, b7) :: b6)) as z. - destruct z. - - comp_simp. - inline_first - inline_first; comp_skip; comp_simp. - - - SearchAbout oc_compMap. - - inline_first. - - - Qed - - . - - Qed. - - unfold Initialize_wLoop_7, Initialize_wLoop_8. - comp_simp. - - - eapply RndPerm_wf. - - eapply (@list_hybrid_close _ _ _ _ _ (G7_2_A2 (a1, b2, b1, b0, b)) (G7_2_A3 (a1, b2, b1, b0, b)) _ _ _ _ (G7_2_A3 (a1, b2, b1, b0, b))). - eapply max_keywords_G7_2_A1. - - intuition. - - unfold LH_Gn. - eapply evalDist_bind_distance. - admit. - intuition. - intuition. - comp_simp. - - destruct (ge_dec n (length a)). - eapply evalDist_bind_distance; intuition. - admit. - - repeat rewrite firstn_ge_all. - intuition. - omega. - omega. - repeat rewrite skipn_length_nil. - eapply leRat_trans. - eapply eqRat_impl_leRat. - rewrite <- ratIdentityIndiscernables. - intuition. - eapply rat0_le_all. - omega. - omega. - - assert ( - evalDist - (lsb1 <-$ (foreach (x in firstn n a)G7_3_A2 (a1, b2, b1, b0, b) x); - lsb2 <-$ (foreach (x in skipn (S n) a)G7_2_A2 (a1, b2, b1, b0, b) x); - x <- nth n a nil; - y <-$ G7_2_A2 (a1, b2, b1, b0, b) x; - G7_2_A3 (a1, b2, b1, b0, b) (lsb1 ++ (y :: lsb2)) - ) a0 == - evalDist - (lsb1 <-$ (foreach (x in firstn n a)G7_3_A2 (a1, b2, b1, b0, b) x); - lsb2 <-$ (foreach (x in skipn n a)G7_2_A2 (a1, b2, b1, b0, b) x); - G7_2_A3 (a1, b2, b1, b0, b) (lsb1 ++ lsb2) - ) a0 - ). - - comp_skip. - destruct a. - simpl in *. - omega. - Local Opaque evalDist. - simpl. - - hwew - - SearchAbout firstn. - comp_skip. - - Qed. - - [db, q, s_A] <-$3 A1; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_8 db k_I k_Z) (toW db); - - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - - Definition Initialize_wLoop_6_1 db k_E k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_6_1 (db : DB)(q : list Query) := - - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - k_Z <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (fun w => k_E <--$ query w; $ Initialize_wLoop_6_1 db k_E k_I k_Z w) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G6_1_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_5 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - - - - Definition G6_1 := - [b, _] <-$2 G5_A _ _ (fun (s : unit)(x : Blist) => r <-$ {0, 1}^lambda; ret (r, s)) tt; - ret b. - - Theorem G6_1_equiv : - Pr[G6] == Pr[G6_1]. - - - unfold G6, G6_1, G5_A. - - eapply comp_spec_eq_impl_eq. - do 4 (simpl; inline_first; comp_skip; comp_simp). - simpl. - inline_first. - comp_skip. - - Theorem oc_compMap_equiv_inv : - - - SearchAbout oc_compMap. - - - Theorem Init_5_o_equiv : - forall a b0, - comp_spec (fun a b => fst a = fst b) - ((Initialize_5 a b0) _ _ (randomFunc ({ 0 , 1 }^lambda) (list_EqDec bool_EqDec)) nil) - ((Initialize_5 a b0) _ _ (fun (s : unit) (_ : Blist) => r <-$ { 0 , 1 }^lambda; ret (r, s)) tt). - - intuition. - simpl; inline_first; comp_skip. - inline_first - - Qed. - - - simpl. - - Qed. - - - Definition Initialize_wLoop_6_1 db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$$ ({0, 1}^lambda); - indLoopRes <--$$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))); - OC_Ret Blist (Bvector lambda) (ret (indLoopRes, sigma)). - - Definition Initialize_6_1 (db : DB)(q : list Query) := - - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - k_Z <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_6_1 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G6_1_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_6_1 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - Definition G6_1 := - [b, _] <-$2 G6_1_A _ _ (@randomFunc Blist (Bvector lambda) (Rnd lambda) _) nil; - ret b. - - Theorem G6_1_equiv : - Pr[G6] == Pr[G6_1]. - - unfold G6, G6_1, G5_A, G6_1_A, Initialize_5, Initialize_6_1. - eapply comp_spec_eq_impl_eq. - simpl; inline_first; comp_skip; comp_simp. - - Qed. - - - - Theorem G4_equiv : - Pr[G3] == Pr[G4]. - - - eapply eq_impl_list_pred. - rewrite H8. - - rewrite H8. - - des - - eapply list_ - SearchAbout length split. - -*) - - (* - assert ( - list_pred - (fun a3 b2 : option (Vector.t bool lambda) => - match a3 with - | Some x3 => match b2 with - | Some y => True - | None => False - end - | None => match b2 with - | Some _ => False - | None => True - end - end) - (ServerSearch_2 (XSetSetup_2 x0 x1 d l3) - (map - (fun c : nat => groupExp g (mult (F_P x0 k0) (F_P x2 (app k c)))) - (allNatsLt (length (fst (split l0))))) (fst (split l0))) - (map - (fun x3 : option (prod (Vector.t bool lambda) (Bvector lambda)) => - match x3 with - | Some v => Some (fst v) - | None => None - end) - (ServerSearch_4 (XSetSetup_2 x0 x1 d l3) - (map - (fun c : nat => - groupExp g (mult (F_P x0 k0) (F_P x2 (app k c)))) - (allNatsLt (length l0))) l0)) - ). -*) - - Unset Printing Notations. - Show. - - f_equal. - eapply map_ext_pred. - eapply list_pred_getSomes. - - SearchAbout length split. - f_equal. - f_equal. - - - Focus 2. - assert (l2 = (fst (split b0))). - rewrite <- Heqz0. - trivial. - subst. - - - - case_eq (arrayLookup (list_EqDec bool_EqDec) (combine (toW d) l2) b1); intuition. - - - - SearchAbout arrayLookupList. - - eapply list_pred_impl. - eapply list_pred_eq. - - rewrite H7. - comp_skip. - inline_first. - - comp_skip. - - comp_skip. - - assert - - Qed. - - - - - Theorem G4_equiv : - Pr[G3] == Pr[G4]. - - unfold G3, G4, G_gen, Initialize_3, Initialize_4. - do 6 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem init_intLoop_4_spec : - forall x1 x2 y b0 z x, - comp_spec (fun a b => fst b = a) - (Initialize_indLoop_2 x1 x2 y b0 (z, x)) - (Initialize_indLoop_4 x1 x2 y b0 (z, x)). - - intuition. - unfold Initialize_indLoop_2, Initialize_indLoop_4. - comp_skip; try - eapply (oneVector (lambda + lambda)). - - eapply comp_spec_ret; intuition. - - Qed. - - Theorem init_wLoop_4_spec : - forall d x x1 x2 b0, - comp_spec (fun a b => snd a = snd b /\ fst a = fst (split (fst b))) - (Initialize_wLoop_2 d x x1 x2 b0) - (Initialize_wLoop_4 d x x1 x2 b0). - - unfold Initialize_wLoop_2, Initialize_wLoop_4. - intuition. - comp_simp. - comp_skip. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - eapply init_intLoop_4_spec. - eapply comp_spec_ret; intuition. - simpl. - split_eq. - eapply list_pred_impl. - eauto. - intuition. - - Qed. - - eapply init_wLoop_4_spec. - - remember (split a1) as z. - destruct z. - remember (split b0) as z. - destruct z. - remember ( split (foreach (ls in l2)split ls)) as z. - destruct z. - comp_skip. - eapply comp_spec_eq_refl_gen. - f_equal. - f_equal. - assert (l0 = fst (split a1)). - rewrite <- Heqz. - trivial. - subst. - assert (l4 = fst (split (foreach (ls in (fst (split b0)))split ls))). - rewrite <- Heqz0. - simpl. - rewrite <- Heqz1. - trivial. - subst. - split_eq. - eapply list_pred_map_r'. - split_eq. - eapply list_pred_impl. - eapply H6. - intuition. - - subst. - assert (l1 = l3). - assert (l1 = snd (split a1)). - rewrite <- Heqz. - trivial. - subst. - assert (l3 = snd (split b0)). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_impl; eauto. - intuition. - subst. - - comp_skip. - - eapply (@compMap_spec _ _ _ _ _ _ eq (fun a b => fst a = fst b /\ snd a = fst (split (snd b)))). - eapply list_pred_eq. - intuition. - subst. - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - eapply comp_spec_ret; intuition. - simpl. - - assert (list_pred (fun a b => a = fst (split b)) l0 l2). - assert (l0 = fst (split a1)). - rewrite <- Heqz. - trivial. - subst. - assert (l2 = fst (split b0)). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eapply list_pred_impl. - eapply H6. - intuition. - - unfold arrayLookupList. - - case_eq ( @arrayLookup Keyword - (list - (Vector.t bool (plus (posnatToNat lambda) (posnatToNat lambda)))) - (@list_EqDec bool bool_EqDec) - (@combine Keyword - (list - (Vector.t bool - (plus (posnatToNat lambda) (posnatToNat lambda)))) - (toW d) l0) b2); intuition. - - - Theorem arrayLookup_Some_list_pred : - forall (A B C : Set){eqda : EqDec A}(P : B -> C -> Prop)(lsa : list A) lsb lsc a b, - list_pred P lsb lsc -> - arrayLookup _ (combine lsa lsb) a = Some b-> - (exists c, - arrayLookup _ (combine lsa lsc) a = Some c /\ - P b c). - - induction lsa; intuition; simpl in *. - discriminate. - destruct lsb. - inversion H; clear H; subst. - simpl in *. - discriminate. - - inversion H; clear H; subst. - simpl in *. - case_eq (eqb a0 a); intuition. - rewrite eqb_leibniz in H. - subst. - rewrite eqb_refl in *. - inversion H0; clear H0; subst. - econstructor; eauto. - - rewrite H in H0. - eapply IHlsa; eauto. - - Qed. - - Theorem arrayLookup_None_combine : - forall (A B C : Set){eqd : EqDec A}(lsa : list A)(lsb : list B)(lsc : list C) a, - arrayLookup _ (combine lsa lsb) a = None -> - length lsb = length lsc -> - arrayLookup _ (combine lsa lsc) a = None. - - induction lsa; intuition; simpl in *. - destruct lsc; simpl in *. - trivial. - destruct lsb; simpl in *. - omega. - destruct ( eqb a0 a). - discriminate. - eauto. - - Qed. - - - specialize (arrayLookup_Some_list_pred _ _ H14 H15); intuition. - destruct H16. - intuition. - rewrite H17. - trivial. - erewrite arrayLookup_None_combine. - trivial. - eauto. - assert (l0 = fst (split a1)). - rewrite <- Heqz. - trivial. - subst. - assert (l2 = fst (split b0)). - rewrite <- Heqz0. - trivial. - subst. - repeat rewrite split_length_l. - eapply list_pred_length_eq; eauto. - - eapply comp_spec_ret; intuition. - f_equal. - - - eapply (@compMap_support _ _ - (fun a b => - map (fun v => Dec (F x a) (fst (splitVector lambda lambda v))) (fst (split (fst b))) = - (snd (split (fst b))) - )) in H5. - - Focus 2. - intuition. - unfold Initialize_wLoop_4 in H14. - repeat simp_in_support. - - eapply (@compMap_support _ _ - (fun a b => Dec (F x a4) (fst (splitVector lambda lambda (fst b))) = snd b)) in H16. - - simpl. - - split_eq. - eapply list_pred_map_l'. - split_eq. - - - - eapply list_pred_impl_1_r. - eapply H16. - intuition. - intuition. - unfold Initialize_indLoop_4 in *. - repeat simp_in_support. - simpl. - - - - rewrite splitVector_append . - simpl. - eapply Enc_correct. - trivial. - - assert (list_pred - (fun (a : Keyword) - (b : list (Vector.t bool (lambda + lambda) * Bvector lambda)) => - (foreach (v in fst (split b)) - Dec (F x a) (fst (splitVector lambda lambda v))) = - snd (split b)) (toW d) l2). - - assert (l2 = (fst (split b0))). - rewrite <- Heqz0. - trivial. - subst. - split_eq. - eauto. - - eapply (@compMap_support _ _ (fun a b => snd b = arrayLookupList _ (combine (toW d) l2) a)) in H11. - Focus 2. - intuition. - repeat simp_in_support. - simpl. - trivial. - - assert (list_pred - (fun a b => fst a = fst b /\ - fst (snd a) = fst (snd b) /\ - snd (snd a) = fst (split (snd (snd b))) /\ - snd (split (snd (snd b))) = - map (fun v => Dec (F x (fst (fst a))) (fst (splitVector lambda lambda v))) (snd (snd a))) - (combine l a2) - (combine l b2) - ). - - admit. - - eapply map_ext_pred. - eauto. - intuition. - simpl in *. - subst. - unfold GenTrans_4. - destruct b3. - simpl. - destruct p0. - simpl in *. - comp_simp. - f_equal. - eapply map_ext_pred. - - - - eapply list_pred_impl. - eapply (@list_pred_combine_l _ _ _ (fun a b => a = fst b /\ - (fun (a : Keyword) - (b : Tag * - list (Vector.t bool (lambda + lambda) * Identifier lambda)) => - snd b = - arrayLookupList (list_EqDec bool_EqDec) (combine (toW d) l2) a) (fst a) (snd b) - ) - (fun a b => (fun (a : Tag * list (Vector.t bool (lambda + lambda))) - (b : Tag * - list (Vector.t bool (lambda + lambda) * Identifier lambda)) => - fst a = fst b /\ snd a = fst (split (snd b))) a (snd b))). - eapply list_pred_symm. - eapply (@list_pred_impl). - eapply (@list_pred_combine_l _ _ _ eq (fun b a => snd b = - arrayLookupList (list_EqDec bool_EqDec) (combine (toW d) l2) (fst a))). - eapply list_pred_eq. - eapply list_pred_symm. - - - - eapply (@list_pred_impl). - eapply list_pred_fst_split_irr_if. - eapply H11. - intuition. - intuition. - - admit. - - intuition. - simpl in *. - subst. - destruct b3. - simpl in *. - subst. - destruct p1. simpl in *. subst. - - - eapply list_pred_I. - - - simpl in *. - eapply H14. - - intros. - eapply H14. - - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_I. - admit. - eapply list_pred_symm. - eapply H12. - intuition. - simpl in *. - subst. - eapply H17. - simpl in *. - intuition. - simpl in *; subst. - simpl. - - - - - SearchAbout list_pred split. - - T - - SearchAbout list_pred combine. - - - admit. - - eapply map_ext_pred. - eauto. - intuition. - simpl in *. - subst. - unfold GenTrans_3, GenTrans_4. - destruct b3. - simpl. - destruct p0. - comp_simp. - f_equal. - - - - simpl in - - eapply list_pred_combine_simple. - eapply list_pred_eq. - eapply H12. - apply compMap_length in H10. - rewrite H10. - symmetry. - eapply split_length_l. - - intuition. - simpl in *. - subst. - unfold GenTrans_4. - destruct b3. - simpl. - destruct p0. - comp_simp. - simpl. - f_equal. - eapply map_ext_pred. - - SearchAbout length split. - - unfold setLet. - inline_first. - comp_simp. - - remember (foreach (ls in l2)split ls) as z. - destruct z. - - remember (split ls) as z. - - Qed. - - - Definition GenTrans_4 (edb : EDB) k_X k_Z e : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t_inds] <-2 e; - [stag, t, inds] <-3 stag_t_inds; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_4 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags_ts_inds <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x), lookupInds db x)) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts_inds); - ret (edb, searchRes). - - Definition G4 := G_gen Initialize_4. - - Theorem G4_equiv : - Pr[G3] == Pr[G4]. - - unfold G3, G4, G_gen, Initialize_3, Initialize_4. - do 8 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply (@compMap_spec _ _ _ _ _ _ eq (fun a b => fst a = fst (fst b) /\ snd a = snd (fst b))). - eapply list_pred_eq. - intuition; subst. - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - eapply comp_spec_ret; intuition. - eapply comp_spec_ret; intuition. - f_equal. - eapply map_ext_pred. - eapply list_pred_combine_simple. - eapply list_pred_eq. - eauto. - eapply compMap_length in H7. - rewrite H7. - symmetry. - eapply split_length_l. - - intuition. - simpl in *. - subst. - unfold GenTrans_4. - destruct b1. - simpl. - comp_simp. - simpl. - f_equal. - - - Qed. - - (* Step 3: For the X, I, and Z PRFs, translate the game so that it interacts with the PRF as an oracle, then replace the PRF with a random function. We only need to replace these three PRFs in order to get to the next step where we remove the decryption. *) - - - (* Step 2.1 : It's a little easier if we turn some map operations into compMaps. SO we do that first. *) - Definition XSetSetup_indLoop_2_1 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - ret (g^(e * xind)). - - Definition XSetSetup_wLoop_2_1 k_X k_I db e := - [w, sigma] <-2 e; - compMap _ (XSetSetup_indLoop_2_1 k_X k_I w) (permute (oneVector lambda) (lookupInds db w) sigma). - - Definition XSetSetup_2_1 k_X k_I db sigmas := - mapRes <-$ compMap _ (XSetSetup_wLoop_2_1 k_X k_I db) (combine (toW db) sigmas); - ret flatten mapRes. - - Definition GenTrans_2_1 (edb : EDB) k_X k_Z k_S (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <- F_P k_X x; - xtokens <-$ compMap _ (fun (c : nat) => ret g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_2_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <-$ XSetSetup_2_1 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <-$ compMap _ (GenTrans_2_1 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - - Theorem G2_1_equiv : - Pr[G2] == Pr[G_gen Initialize_2_1]. - - unfold G2, G_gen, Initialize_2, Initialize_2_1. - - do 7 (comp_skip; comp_simp). - comp_irr_r. - - Theorem compMap_wf : - forall (A B : Set)(eqdb : EqDec B)(ls : list A)(c : A -> Comp B), - (forall a, In a ls -> well_formed_comp (c a)) -> - well_formed_comp (compMap _ c ls). - - induction ls; intuition; simpl in *; wftac. - - Qed. - - Hint Resolve compMap_wf : wftac. - - unfold XSetSetup_2_1; wftac. - eapply compMap_wf; intuition. - unfold XSetSetup_wLoop_2_1. - wftac. - eapply compMap_wf. - intuition. - unfold XSetSetup_indLoop_2_1. - wftac. - - comp_skip. - comp_simp. - comp_skip. - - comp_irr_r. - eapply compMap_wf. - intuition. - unfold GenTrans_2_1. - wftac. - eapply compMap_wf. - intuition. - wftac. - - eapply evalDist_ret_eq. - - assert ( XSetSetup_2 x0 x1 d l1 = x4). - - unfold XSetSetup_2_1 in *. - repeat simp_in_support. - unfold XSetSetup_2. - f_equal. - - eapply (@compMap_support _ _ (fun (a : Blist * (list nat)) b => b = XSetSetup_wLoop_2 x0 x1 d a)) in H9. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply H9. - intuition. - intuition. - unfold XSetSetup_wLoop_2_1 in *. - eapply (@compMap_support _ _ (fun a b => b = XSetSetup_indLoop_2 x0 x1 a2 a)) in H10. - unfold XSetSetup_wLoop_2. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_r'. - eapply list_pred_symm. - trivial. - intuition. - unfold XSetSetup_indLoop_2_1, XSetSetup_indLoop_2 in *. - repeat simp_in_support. - trivial. - - subst. - f_equal. - - apply (@compMap_support _ _ (fun a b => b = GenTrans_2 (t, XSetSetup_2 x0 x1 d l1) x0 x2 x a)) in H8. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply H8. - intuition. - intuition. - unfold GenTrans_2, GenTrans_2_1 in *. - comp_simp. - repeat simp_in_support. - - eapply (@compMap_support _ _ (fun (a : nat) b => b = g ^ (F_P x0 k0 * F_P x2 (k ++ a)))) in H11. - - assert (x4 = map (fun (a : nat) => g ^ (F_P x0 k0 * F_P x2 (k ++ a))) (allNatsLt (length (TSetRetrieve t b0)))). - eapply list_pred_eq_impl_eq. - eapply list_pred_map_r'. - eapply list_pred_symm. - trivial. - subst. - - trivial. - - intuition. - repeat simp_in_support. - trivial. - - Qed. - - - (* Step 2.2: We will need to modify the top-level game to apply the PRF definition. Start by putting everything in one function. *) - Definition G2_2 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <-$ XSetSetup_2_1 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <-$ compMap _ (GenTrans_2_1 edb k_X k_Z k_S) (combine q stags); - A2 s_A edb (snd (split searchRes)). - - Theorem G2_2_equiv : - Pr[G_gen Initialize_2_1] == Pr[G2_2]. - - unfold G_gen, Initialize_2_1, G2_2. - do 10 (comp_skip; comp_simp; inline_first). - reflexivity. - Qed. - - (* Step 2.3: replace the "X" PRF first. Start by sampling the key at the top of beginning of the game.*) - - Definition G2_3 := - k_X <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <-$ XSetSetup_2_1 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <-$ compMap _ (GenTrans_2_1 edb k_X k_Z k_S) (combine q stags); - A2 s_A edb (snd (split searchRes)). - - Theorem G2_3_equiv : - Pr[G2_2] == Pr[G2_3]. - - unfold G2_2, G2_3. - do 2 (comp_swap_r; comp_skip; comp_simp). - - Qed. - - (* Step 2.4 : Put the game in the form of the PRF definition. *) - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - Notation "'query' v" := (OC_Query _ v) (at level 79) : comp_scope. - - Definition XSetSetup_indLoop_2_4 k_I (w : Blist) (ind : Bvector lambda) := - e <--$ query w; - xind <- F_P k_I (Vector.to_list ind); - $ ret (g^(e * xind)). - - Definition XSetSetup_wLoop_2_4 k_I db e := - [w, sigma] <-2 e; - oc_compMap _ (XSetSetup_indLoop_2_4 k_I w) (permute (oneVector lambda) (lookupInds db w) sigma). - - Definition XSetSetup_2_4 k_I db sigmas := - mapRes <--$ oc_compMap _ (XSetSetup_wLoop_2_4 k_I db) (combine (toW db) sigmas); - $ ret flatten mapRes. - - Definition GenTrans_2_4 (edb : EDB) k_Z k_S (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <--$ query x; - xtokens <--$$ compMap _ (fun (c : nat) => ret g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - $ ret (inds, (stag, (combine xtokens res))). - - Definition PRF_A_1 := - [db, q, s_A] <--$3$ A1; - k_S <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - k_Z <--$$ {0,1}^lambda; - wLoopRes <--$$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <--$ XSetSetup_2_4 k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - edb <- (tSet, xSet); - stags <--$$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_2_4 edb k_Z k_S) (combine q stags); - $ A2 s_A edb (snd (split searchRes)). - - Definition G2_4 := - k_X <-$ {0,1}^lambda; - [r, _] <-$2 (PRF_A_1 _ _ (f_oracle F_P _ k_X) tt); - ret r. - - - Theorem compMap_oc_spec : - forall (C D : Set)(P2 : C -> D -> Prop)(A B : Set)(P1 : A -> B -> Prop)(eqdc : EqDec C)(eqdd : EqDec D)(E F S: Set)(eqds : EqDec S)(ls1 : list A)(ls2 : list B)(c1 : A -> Comp C)(c2 : B -> OracleComp E F D)o (s : S), - list_pred P1 ls1 ls2 -> - (forall a b z, P1 a b -> comp_spec (fun x y => P2 x (fst y)) (c1 a) (c2 b _ _ o z)) -> - comp_spec (fun a b => list_pred P2 a (fst b)) - (compMap _ c1 ls1) - ((oc_compMap _ c2 ls2) _ _ o s). - - induction ls1; inversion 1; subst; intuition; simpl in *. - comp_simp. - eapply comp_spec_ret; simpl; econstructor. - - simpl. - comp_skip. - comp_simp. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - econstructor; eauto. - - Qed. - - Theorem G2_4_equiv : - Pr[G2_3] == Pr[G2_4]. - - unfold G2_3, G2_4, PRF_A_1. - eapply comp_spec_eq_impl_eq. - simpl. - (simpl; inline_first; comp_skip; comp_simp). - simpl; inline_first. - comp_skip. - comp_simp. - - do 3 (simpl; inline_first; comp_skip; comp_simp). - - unfold XSetSetup_2_1, XSetSetup_2_4. - simpl; inline_first. - comp_skip. - comp_ret_r. - inline_first. - comp_simp. - simpl. - inline_first. - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold XSetSetup_wLoop_2_1, XSetSetup_wLoop_2_4. - eapply comp_spec_consequence. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - unfold XSetSetup_indLoop_2_1, XSetSetup_indLoop_2_4. - simpl. - unfold f_oracle. - comp_simp. - eapply comp_spec_ret; intuition. - intuition. - eapply list_pred_eq_impl_eq. - trivial. - - comp_simp. - inline_first. - comp_ret_r. - comp_ret_r. - comp_simp. - inline_first. - comp_skip. - comp_ret_r. - comp_simp. - simpl. - inline_first. - comp_skip. - comp_ret_r. - comp_simp. - inline_first. - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold GenTrans_2_1, GenTrans_2_4. - comp_simp. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - repeat f_equal. - eapply list_pred_eq_impl_eq; trivial. - eapply list_pred_eq_impl_eq; trivial. - - comp_simp. - inline_first. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - assert (a0 = l1). - eapply list_pred_eq_impl_eq. - trivial. - subst. - assert (a1 = l2). - eapply list_pred_eq_impl_eq. - trivial. - subst. - eapply comp_spec_eq_refl. - subst. - comp_simp. - eapply comp_spec_eq_refl. - - Qed. - - (* Step 2.5: replace the PRF with a random function. *) - - Definition G2_5 := - [r, _] <-$2 (PRF_A_1 _ _ (@randomFunc Blist nat (RndG) _) nil); - ret r. - - Theorem G2_5_equiv : - | Pr[G2_4] - Pr[G2_5] | == PRF_Advantage (Rnd lambda) (RndG) F_P _ _ PRF_A_1. - - reflexivity. - - Qed. - - - (* Step 3: prepare to replace all PRFs with random functions. We use a composite oracle to manage the 4 PRFs. *) - - Notation "'query' v" := (OC_Query _ v) (at level 79) : comp_scope. - - Instance sum_EqDec : - forall (A B : Set)(eqda : EqDec A)(eqdb : EqDec B), - EqDec (A + B). - - intuition. - exists - (fun x y => - match x with - | inl a => - match y with - | inl a' => eqb a a' - | inr _ => false - end - | inr b => - match y with - | inl _ => false - | inr b' => eqb b b' - end - end). - - intuition; try discriminate. - f_equal. - eapply eqb_leibniz. - trivial. - inversion H; clear H; subst. - eapply eqb_leibniz. - trivial. - - f_equal. - eapply eqb_leibniz. - trivial. - inversion H; clear H; subst. - eapply eqb_leibniz. - trivial. - Qed. - - Definition mkCompositeOracle - {A1 A2 B1 B2 S1 S2 : Set} - {eqdb1 : EqDec B1}{eqdb2 : EqDec B2} - {eqds1 : EqDec S1}{eqds2 : EqDec S2} - (o1 : S1 -> A1 -> Comp (B1 * S1)) - (o2 : S2 -> A2 -> Comp (B2 * S2)) - (s : S1 * S2)(z : A1 + A2) := - match z with - | inl a => [b, s'] <-$2 o1 (fst s) a; ret (inl b, (s', snd s)) - | inr a => [b, s'] <-$2 o2 (snd s) a; ret (inr b, (fst s, s')) - end. - - Definition toL(A B : Type)(def : A)(x : A + B) : A := - match x with - | inl a => a - | inr _ => def - end. - - Definition toR(A B : Type)(def : B)(x : A + B) : B := - match x with - | inl _ => def - | inr b => b - end. - - - (* - / \ - / \ / \ - S X I Z *) - - Definition D_O := (Blist + Blist + (Blist + Blist))%type. - Definition R_O := (Bvector lambda + nat + (nat + nat))%type. - - Definition QS (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QX (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - Definition QI (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QZ (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - - Definition toS(r : R_O) := - x <- toL (inl nat (oneVector lambda)) r; - toL (oneVector lambda) x. - - Definition toX(r : R_O) := - x <- toL (inl nat (oneVector lambda)) r; - toR O x. - - Definition toI(r : R_O) := - x <- toR (inl nat O) r; - toL O x. - - Definition toZ(r : R_O) := - x <- toR (inl nat O) r; - toR O x. - - Definition sQuery(x : Blist) := - qRes <--$ query (QS x); - $ ret (toS qRes). - - Definition xQuery(x : Blist) := - qRes <--$ query (QX x); - $ ret (toX qRes). - - Definition iQuery(x : Blist) := - qRes <--$ query (QI x); - $ ret (toI qRes). - - Definition zQuery(x : Blist) := - qRes <--$ query (QZ x); - $ ret (toZ qRes). - - - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - Definition Initialize_indLoop_3 k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E ind; - xind <--$ iQuery (Vector.to_list ind); - z <--$ zQuery (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y)). - - Definition Initialize_wLoop_3 db w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$ sQuery w; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_3 k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition XSetSetup_indLoop_3 w (ind : Bvector lambda) := - e <--$ xQuery w; - xind <--$ iQuery (Vector.to_list ind); - $ ret g^(e * xind). - - Definition XSetSetup_wLoop_3 db e := - [w, sigma] <-2 e; - oc_compMap _ (XSetSetup_indLoop_3 w) (permute (oneVector lambda) (lookupInds db w) sigma). - - Definition XSetSetup_3 db sigmas := - loopRes <--$ oc_compMap _ (XSetSetup_wLoop_3 db) (combine (toW db) sigmas); - $ ret flatten loopRes. - - - Definition ServerSearch_3 := ServerSearch_2. - - Definition GenTrans_3 (edb : EDB) (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <--$ xQuery x; - xtokens <--$ oc_compMap _ (fun (c : nat) => exp <--$ zQuery (s ++ c); $ ret g^(e * exp)) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <--$ sQuery s; - inds <- map (Dec k_E) es; - $ ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_3_oc (db : DB)(q : list Query) := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_3 db) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <--$ XSetSetup_3 db sigmas; - t <- combine (toW db) t_entries; - tSetRes <--$$ TSetSetup t; - [tSet, k_T] <-2 tSetRes; - edb <- (tSet, xSet); - stags <--$$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_3 edb) (combine q stags); - $ ret (edb, searchRes). - - Definition Initialize_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o1 := mkCompositeOracle (f_oracle F _ k_S) (f_oracle F_P _ k_X) in - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_3_oc db q) _ _ o (tt, tt, (tt, tt)); - ret r. - - Definition G3 := G_gen Initialize_3. - - - (* Step 2.1 : It's a little easier if we turn some map operations into compMaps. *) - Definition XSetSetup_indLoop_2_1 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - ret (g^(e * xind)). - - Definition XSetSetup_wLoop_2_1 k_X k_I db e := - [w, sigma] <-2 e; - compMap _ (XSetSetup_indLoop_2_1 k_X k_I w) (permute (oneVector lambda) (lookupInds db w) sigma). - - Definition XSetSetup_2_1 k_X k_I db sigmas := - mapRes <-$ compMap _ (XSetSetup_wLoop_2_1 k_X k_I db) (combine (toW db) sigmas); - ret flatten mapRes. - - Definition GenTrans_2_1 (edb : EDB) k_X k_Z k_S (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <- F_P k_X x; - xtokens <-$ compMap _ (fun (c : nat) => ret g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_2_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <-$ XSetSetup_2_1 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <-$ compMap _ (GenTrans_2_1 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - Theorem G2_1_equiv : - Pr[G2] == Pr[G_gen Initialize_2_1]. - - unfold G2, G_gen, Initialize_2, Initialize_2_1. - - do 7 (comp_skip; comp_simp). - comp_irr_r. - - Theorem compMap_wf : - forall (A B : Set)(eqdb : EqDec B)(ls : list A)(c : A -> Comp B), - (forall a, In a ls -> well_formed_comp (c a)) -> - well_formed_comp (compMap _ c ls). - - induction ls; intuition; simpl in *; wftac. - - Qed. - - Hint Resolve compMap_wf : wftac. - - unfold XSetSetup_2_1; wftac. - eapply compMap_wf; intuition. - unfold XSetSetup_wLoop_2_1. - wftac. - eapply compMap_wf. - intuition. - unfold XSetSetup_indLoop_2_1. - wftac. - - comp_skip. - comp_simp. - comp_skip. - - comp_irr_r. - eapply compMap_wf. - intuition. - unfold GenTrans_2_1. - wftac. - eapply compMap_wf. - intuition. - wftac. - - eapply evalDist_ret_eq. - - assert ( XSetSetup_2 x0 x1 d l1 = x4). - - unfold XSetSetup_2_1 in *. - repeat simp_in_support. - unfold XSetSetup_2. - f_equal. - - eapply (@compMap_support _ _ (fun (a : Blist * (list nat)) b => b = XSetSetup_wLoop_2 x0 x1 d a)) in H9. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply H9. - intuition. - intuition. - unfold XSetSetup_wLoop_2_1 in *. - eapply (@compMap_support _ _ (fun a b => b = XSetSetup_indLoop_2 x0 x1 a2 a)) in H10. - unfold XSetSetup_wLoop_2. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_r'. - eapply list_pred_symm. - trivial. - intuition. - unfold XSetSetup_indLoop_2_1, XSetSetup_indLoop_2 in *. - repeat simp_in_support. - trivial. - - subst. - f_equal. - - apply (@compMap_support _ _ (fun a b => b = GenTrans_2 (t, XSetSetup_2 x0 x1 d l1) x0 x2 x a)) in H8. - eapply list_pred_eq_impl_eq. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply H8. - intuition. - intuition. - unfold GenTrans_2, GenTrans_2_1 in *. - comp_simp. - repeat simp_in_support. - - eapply (@compMap_support _ _ (fun (a : nat) b => b = g ^ (F_P x0 k0 * F_P x2 (k ++ a)))) in H11. - - assert (x4 = map (fun (a : nat) => g ^ (F_P x0 k0 * F_P x2 (k ++ a))) (allNatsLt (length (TSetRetrieve t b0)))). - eapply list_pred_eq_impl_eq. - eapply list_pred_map_r'. - eapply list_pred_symm. - trivial. - subst. - - trivial. - - intuition. - repeat simp_in_support. - trivial. - - Qed. - - Theorem G3_equiv : - Pr[G_gen Initialize_2_1] == Pr[G3]. - - unfold G3, G_gen, Initialize_2_1, Initialize_3, Initialize_3_oc. - - do 6 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - simpl. - inline_first. - - eapply (@comp_spec_seq _ _ (fun a b => a = fst b)); eauto with inhabited. - - Theorem compMap_oc_spec : - forall (C D : Set)(P2 : C -> D -> Prop)(A B : Set)(P1 : A -> B -> Prop)(eqdc : EqDec C)(eqdd : EqDec D)(E F S: Set)(eqds : EqDec S)(ls1 : list A)(ls2 : list B)(c1 : A -> Comp C)(c2 : B -> OracleComp E F D)o (s : S), - list_pred P1 ls1 ls2 -> - (forall a b z, P1 a b -> comp_spec (fun x y => P2 x (fst y)) (c1 a) (c2 b _ _ o z)) -> - comp_spec (fun a b => list_pred P2 a (fst b)) - (compMap _ c1 ls1) - ((oc_compMap _ c2 ls2) _ _ o s). - - induction ls1; inversion 1; subst; intuition; simpl in *. - comp_simp. - eapply comp_spec_ret; simpl; econstructor. - - simpl. - comp_skip. - comp_simp. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - econstructor; eauto. - - Qed. - - eapply (comp_spec_consequence (fun a b => a = fst b)). - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - unfold Initialize_wLoop_2, Initialize_wLoop_3. - simpl. - comp_simp. - inline_first. - comp_skip. - - (* comp_simp gets a little slow here *) - comp_ret_r. - inline_first. - unfold f_oracle. - comp_ret_r. - comp_simp. - unfold toS. - inline_first. - unfold toL. - comp_ret_r. - comp_ret_r. - simpl. - - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold Initialize_indLoop_2, Initialize_indLoop_3. - simpl. - inline_first. - comp_skip; try eapply (oneVector (lambda + lambda)). - - repeat (try comp_ret_r; inline_first). - eapply comp_spec_ret. - simpl. - unfold toI. - comp_simp. - unfold toZ, toL, toR. - comp_simp. - trivial. - - comp_simp. - eapply comp_spec_ret; intuition. - simpl in *. - f_equal. - eapply list_pred_eq_impl_eq. - trivial. - - intuition. - eapply list_pred_eq_impl_eq; trivial. - - intuition. - subst. - destruct b0. - simpl. - comp_simp. - simpl. - unfold XSetSetup_2_1. - inline_first. - - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - unfold XSetSetup_wLoop_2_1, XSetSetup_wLoop_3. - eapply comp_spec_consequence. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - unfold XSetSetup_indLoop_2_1, XSetSetup_indLoop_3. - simpl. - unfold f_oracle. - inline_first. - repeat (comp_ret_r; inline_first). - simpl. - unfold toX, toI. - comp_simp. - unfold toR, toL. - eapply comp_spec_ret; intuition. - - intuition. - eapply list_pred_eq_impl_eq. - trivial. - - comp_simp. - inline_first. - repeat comp_ret_r. - comp_simp. - inline_first. - comp_skip. - comp_ret_r. - comp_simp. - simpl. - inline_first. - comp_skip. - comp_ret_r. - comp_simp. - inline_first. - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - pairInv. - - unfold GenTrans_2_1, GenTrans_3. - inline_first. - destruct a2. - simpl. - inline_first. - unfold f_oracle. - repeat (comp_ret_r; inline_first). - comp_skip. - eapply (@compMap_oc_spec _ _ eq). - eapply list_pred_eq. - intuition; subst. - simpl. - inline_first. - repeat (comp_ret_r; inline_first). - unfold toX, toZ. - comp_simp. - unfold toR, toL. - eapply comp_spec_ret; intuition. - - destruct b2. - inline_first. - repeat (comp_ret_r; inline_first). - eapply comp_spec_ret; intuition. - simpl in *. - f_equal. - f_equal. - f_equal. - f_equal. - f_equal. - eapply list_pred_eq_impl_eq. - trivial. - eapply list_pred_eq_impl_eq. - trivial. - - f_equal. - f_equal. - eapply list_pred_eq_impl_eq. - trivial. - f_equal. - f_equal. - eapply list_pred_eq_impl_eq. - trivial. - eapply list_pred_eq_impl_eq. - trivial. - - destruct b2. - inline_first. - repeat comp_ret_r. - eapply comp_spec_ret; intuition. - f_equal. - f_equal. - f_equal. - eapply list_pred_eq_impl_eq. - trivial. - eapply list_pred_eq_impl_eq. - trivial. - Qed. - - Check F_P. - - - - (* Step 4: Replace PRFs with random functions. *) - Definition Initialize_4 (db : DB)(q : list Query) := - - let o1 := mkCompositeOracle - (@randomFunc Keyword (Bvector lambda) (Rnd lambda) _ ) - (@randomFunc Blist nat (RndNat lambda) _) in - let o2 := mkCompositeOracle - (@randomFunc Blist nat (RndNat lambda) _) - (@randomFunc Blist nat (RndNat lambda) _) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_3_oc db q) _ _ o (nil, nil, (nil, nil)); - ret r. - - Definition G4 := G_gen Initialize_4. - - Print OracleComp. - - Definition AWO_OracleClose_r (A B C D E : Set){eqda : EqDec A}{eqdb : EqDec B}{eqdc : EqDec C}{eqdd : EqDec D} - (c : OracleComp (A + B) (C + D) E) - (S : Set)(eqds : EqDec S)(o : S -> B -> Comp (D * S))(s : S) : OracleComp A C (E * S) := - OC_Run _ _ _ c - (fun s p => - match p with - | inl a => (z <--$ query a; $ ret (inl z, s)) - | inr b => [d, s'] <--$2 $ o s b; $ ret (inr d, s') - end) s. - - - (* Replace one oracle at a time, in the order the keys are selected *) - (* Step 3.1: We need to modify the top-level game to unify with the PRF definition. Start by putting everything in one function. *) - - Definition G3_1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o1 := mkCompositeOracle (f_oracle F _ k_S) (f_oracle F_P _ k_X) in - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_3_oc db q) _ _ o (tt, tt, (tt, tt)); - [edb, ls] <-2 r; - A2 s_A edb (snd (split ls)). - - - Definition Initialize_3_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - let p := ( - - )in - [r, _] <-$2 p _ _ (f_oracle F _ k_S) tt; - ret r. - - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o1 := mkCompositeOracle (f_oracle F _ k_S) (f_oracle F_P _ k_X) in - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let p := AWO_OracleClose_r (Initialize_3_oc db q) _ o2 (tt, tt) in - [r, _] <-$2 p _ _ o1 (tt, tt); - ret (fst r). - - Definition Initialize_3_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - - let p := AWO_OracleClose_r (Initialize_3_oc db q) _ o2 (tt, tt) in - let p := AWO_OracleClose_r p _ (f_oracle F_P _ k_X) tt in - [r, _] <-$2 p _ _ (f_oracle F _ k_S) tt; - ret (fst (fst r)). - - Theorem G_3_1_equiv : - Pr[G3] == Pr[G_gen Initialize_3_1]. - - unfold G3, G_gen, Initialize_3, Initialize_3_1. - do 6 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - eapply (@comp_spec_seq _ _ (fun a b => fst a = fst (fst b))); - eauto with inhabited. - - Theorem AWO_OracleClose_r_composite_equiv : - forall (A B C D E S1 S2: Set){eqda : EqDec A}{eqdb : EqDec B}{eqdc : EqDec C}{eqdd : EqDec D}{eqde : EqDec E}{eqds1 : EqDec S1}{eqds2 : EqDec S2}(c : OracleComp (A + B) (C + D) E) - (o1 : S1 -> B -> Comp (D * S1))(o2 : S2 -> A -> Comp (C * S2)) - (s1 : S1) (s2 : S2), - comp_spec (fun p1 p2 => (fst (fst p1) = fst p2 /\ - snd (fst p1) = snd (snd p2) /\ - snd p1 = fst (snd p2))) - ((AWO_OracleClose_r c _ o1 s1) _ _ o2 s2) - (c _ _ (mkCompositeOracle o2 o1) (s2, s1)). - - intuition. - unfold AWO_OracleClose_r, mkCompositeOracle. - simpl. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - eapply oc_base_exists; eauto; intuition; - eapply comp_base_exists in X1; - intuition. - eapply oc_base_exists; eauto; intuition; - eapply comp_base_exists in X1; - intuition. - - eapply (@oc_comp_spec_eq _ _ _ _ _ _ _ _ _ _ _ _ _ _ (fun p1 p2 => (snd p1, fst p1) = p2)). - trivial. - intros. - subst. - simpl. - intuition. - simpl. - inline_first. - comp_skip. - apply comp_base_exists in X1; intuition. - apply comp_base_exists in X1; intuition. - inline_first. - simpl. - eapply comp_spec_ret; intuition. - - simpl. - inline_first. - comp_skip. - apply comp_base_exists in X1; intuition. - apply comp_base_exists in X1; intuition. - simpl. - inline_first; comp_simp. - eapply comp_spec_ret; intuition. - - eapply comp_spec_ret; simpl in *; intuition. - subst. - destruct b; simpl in *; intuition. - destruct p0; simpl in *; intuition. - pairInv; intuition. - - subst. - destruct b; simpl in *. - destruct p0; simpl in *. - pairInv; intuition. - - Qed. - - - eapply comp_spec_symm. - eapply comp_spec_consequence. - eapply AWO_OracleClose_r_composite_equiv. - intuition. - intuition. - simpl in *. - comp_simp. - simpl in *. - eapply comp_spec_ret; intuition. - Qed. - - Theorem G_3_2_equiv : - Pr[G_gen Initialize_3_1] == Pr[G_gen Initialize_3_2]. - - unfold G_gen, Initialize_3_1, Initialize_3_2. - do 6 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply comp_spec_symm. - eapply AWO_OracleClose_r_composite_equiv. - simpl in H6. - intuition. - subst. - comp_simp. - simpl in H7. - eapply comp_spec_ret; intuition. - destruct p0; simpl; simpl in H7. - destruct p0; simpl; simpl in H7. - pairInv. - trivial. - - Qed. - - - (*Step 3.3: move things around and get the game to match the PRF definition. *) - Definition Initialize_3_3 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let p := AWO_OracleClose_r (Initialize_3_oc db q) _ o2 (tt, tt) in - let p := AWO_OracleClose_r p _ (f_oracle F_P _ k_X) tt in - r <-$ - (k_S <-$ {0,1}^lambda; - ([r, _] <-$2 p _ _ (f_oracle F _ k_S) tt; - ret r)); - ret (fst (fst r)). - - - Theorem G_3_3_equiv : - Pr[G_gen Initialize_3_2] == Pr[G_gen Initialize_3_3]. - - unfold G_gen, Initialize_3_2, Initialize_3_3. - do 3 (comp_skip; comp_simp; inline_first; comp_at comp_inline leftc 1%nat; comp_swap_l). - comp_skip. - inline_first. - comp_skip. - comp_inline_l. - comp_inline_r. - comp_skip. - comp_simp. - intuition. - Qed. - - (* Step 3.4 apply the PRF definition *) - Definition Initialize_3_4 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - - let p := AWO_OracleClose_r (Initialize_3_oc db q) _ o2 (tt, tt) in - let p := AWO_OracleClose_r p _ (f_oracle F_P _ k_X) tt in - r <-$ - ([r, _] <-$2 p _ _ (@randomFunc Blist (Bvector lambda) (Rnd lambda) _) nil; - ret r); - ret (fst (fst r)). - - Theorem G3_4_equiv : - | Pr[G_gen Initialize_3_3] - Pr[G_gen Initialize_3_4] | == PRF_Advantange. - Qed. - - Theorem G3_3_equiv : - Pr[G_gen Initialize_3_2] == Pr[G_gen Initialize_3_3]. - - unfold G_gen, Initialize_3_2, Initialize_3_3. - (comp_skip; comp_simp; inline_first). - comp_skip. - - Qed. - - - - Definition Initialize_3_2 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - - let p := AWO_OracleClose_r (Initialize_3_oc db q) _ o2 (tt, tt) in - let p := AWO_OracleClose_r p _ (f_oracle F_P _ k_X) tt in - [r, _] <-$2 p _ _ (@randomFunc Keyword (Bvector lambda) (Rnd lambda) _ ) nil; - ret (fst (fst r)). - - - - unfold toZ. - unfold toR. - - comp_ret_r. - comp_simp. - - Definition Initialize_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - - let s_oracle := f_oracle F _ k_S in - let x_oracle := f_oracle F_P _ k_X in - let i_oracle := f_oracle F _ k_I in - let z_oracle := f_oracle F _ k_Z in - - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_2 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - Definition G3 := G_gen Initialize_3. - - - We made it to G2! continue here - Next step: replace all PRFs with random functions - - (* Step 1.7: Simplify*) - - Definition Initialize_1_7 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet_raw <- map (XSetSetup_wLoop_2 k_X k_I db) (combine (toW db) sigmas); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - stags <-$ compMap _ (TSetGetTag k_T) q - searchRes <-$ map (GenTrans_2 edb k_X k_Z k_S) ls)); - ret (edb, searchRes). - - Definition Initialize_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_2 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - - - subst. - - - eapply list_pred_impl. - eapply (@list_pred_combine_l _ _ _ (fun a b => a = fst b)). - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_eq. - eapply list_pred_I. - admit. - intuition. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - do 2 (split_subst; split_eq). - eapply list_pred_impl. - eapply H4. - intuition. - - - comp_skip. - - - eapply Init_wLoop_1_4_spec. - remember (split a1) as z. - destruct z. - remember (split b0) as z. - destruct z. - assert (l0 = l2). - split_subst. - symmetry. - split_subst. - symmetry. - split_eq. - trivial. - - subst. - comp_skip. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - eapply comp_spec_eq_refl_gen. - f_equal. - f_equal. - f_equal. - - - - assert (l1 = map (fun p => permute (oneVector lambda) (lookupInds d (fst p)) (snd p)) (combine (toW d) l3)). - split_subst. - split_eq. - - - - - - eapply (@map_ext_pred _ _ _ (fun a b => fst a = fst b /\ snd a = permute (oneVector lambda) (lookupInds d (fst a)) (snd b))). - - admit. - intuition. - simpl in*; subst. - unfold XSetSetup_wLoop_1_4. - comp_simp. - simpl. - intuition. - - eapply comp_spec_ret; intuition. - - f_equal. - f_equal. - - - eapply compMap_eq. - - eapply list_pred_eq_impl_eq. - eapply list_pred_map_both'. - - unfold XSetSetup_wLoop_1_3. - - assert (list_pred - (fun (a : Blist * list (Bvector lambda)) (b : Blist * list nat) => - (fst a = fst b) /\ snd a = (permute - - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_symm. - eapply list_pred_impl. - eapply list_pred_combine_l. - eapply list_pred_eq. - eapply list_pred_I. - - eapply list_pre - eapply list_pred_combine_l. - - eapply map_ext. - - Qed. - - (* Now we change how we deal with permutations to get game 2 *) - Theorem G2_equiv : - Pr[G_gen Initialize_1_3] == Pr[G2]. - - unfold G2, G_gen, Initialize_1_3, Initialize_2. - do 6 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_2_spec : forall d x x1 x2 b0, - comp_spec (fun a b => fst a = fst b) - (Initialize_wLoop_1_3 d x x1 x2 b0) - (Initialize_wLoop_2 d x x1 x2 b0). - - intuition. - unfold Initialize_wLoop_1_3, Initialize_wLoop_2. - comp_simp. - comp_skip. - comp_skip. - eapply compMap_spec. - - rewrite permute_length_eq. - apply RndPerm_In_support_length in H. - rewrite H. - eapply list_pred_eq. - intuition; subst. - (*TODO: remove one of these procedures *) - eapply comp_spec_eq_refl. - eapply comp_spec_ret; intuition. - simpl. - eapply list_pred_eq_impl_eq. - trivial. - - Qed. - - eapply Init_wLoop_2_spec. - - remember (split a1) as z. - destruct z. - remember (split b0) as z. - destruct z. - assert (l0 = l2). - split_subst. - symmetry. - split_subst. - symmetry. - split_eq. - trivial. - subst. - comp_skip. - eapply - - intuition. - Qed. - - (* Step 1.4: Change the setup routine so it returns the permutations rather than the indices *) - Definition Initialize_wLoop_1_4 db k_S k_I k_Z w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute (oneVector lambda) inds sigma; - indLoopRes <-$ compMap _ (Initialize_indLoop_1_1 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition XSetSetup_wLoop_1_4 k_X k_I e := - [w, sigma] <-2 e; - map (XSetSetup_indLoop_1_1 k_X k_I w) (permute _ inds sigma). - - Definition Initialize_1_4 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_4 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet_raw <- map (XSetSetup_wLoop_1_3 k_X k_I) (combine (toW db) sigmas); - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - - (* Step 1.1: have the init routine return the permutations used *) - Definition Initialize_wLoop_1_1 db k_S k_I k_Z k_X w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute (oneVector lambda) inds sigma; - indLoopRes <-$ compMap _ (OXT_EDBSetup_indLoop k_I k_Z k_E k_X w) (combine inds (allNatsLt (length inds))); - ret (split indLoopRes, sigma). - - Definition Initialize_1_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_1 db k_S k_I k_Z k_X) (toW db); - [wEntries, perms] <-2 split wLoopRes; - [t_entries, xSet_raw] <-2 split wEntries; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - - - - Theorem G1_1_equiv : - Pr[G_gen Initialize_1] == Pr[G_gen Initialize_1_1]. - - unfold G_gen, Initialize_1, Initialize_1_1. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_1_1_spec : - forall d x x1 x2 x0 b0, - comp_spec - (fun x y => - x = fst y) - (OXT_EDBSetup_wLoop d x x1 x2 x0 b0) - (Initialize_wLoop_1_1 d x x1 x2 x0 b0). - - intuition. - unfold OXT_EDBSetup_wLoop. - unfold Initialize_wLoop_1_1. - comp_simp. - unfold shuffle. - inline_first; comp_skip. - comp_skip. - eapply comp_spec_ret; intuition. - - Qed. - - eapply Init_wLoop_1_1_spec. - intuition. - remember (split b0) as z. - destruct z. - - assert (fst (split b0) = a1). - eapply fst_split_eq_list_pred. - eapply list_pred_symm. - eapply list_pred_impl. - eauto. - intuition. - rewrite <- Heqz in H7. - simpl in H7. - subst. - comp_simp. - eapply comp_spec_eq_refl. - Qed. - - (* Step 1.2: Split off the XSet computation *) - Definition Initialize_wLoop_1_2 db k_S k_I k_Z k_X w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute (oneVector lambda) inds sigma; - indLoopRes <-$ compMap _ (OXT_EDBSetup_indLoop k_I k_Z k_E k_X w) (combine inds (allNatsLt (length inds))); - ret (split indLoopRes, sigma). - - Definition Initialize_1_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_2 db k_S k_I k_Z k_X) (toW db); - [wEntries, perms] <-2 split wLoopRes; - [t_entries, xSet_raw] <-2 split wEntries; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - - Definition OXT_EDBSetup_indLoop_1_2 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - e <-$ Enc k_E ind; - ret (Vector.append e (natToBvector y)). - - Definition Initialize_wLoop_1_2 db k_S k_I k_Z w := - k_E <- F k_S w; - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - inds <- permute (oneVector lambda) inds sigma; - indLoopRes <-$ compMap _ (OXT_EDBSetup_indLoop_1_2 k_I k_Z k_E w) (combine inds (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition XSetSetup_indLoop_1_2 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - g^(e * xind). - - Definition XSetSetup_wLoop_1_2 k_X k_I db e := - [w, sigma] <-2 e; - map (XSetSetup_indLoop_1_2 k_X k_I w) (permute (oneVector lambda) (lookupInds db w) sigma). - - Definition XSetSetup_1_2 k_X k_I db perms := - flatten (map (XSetSetup_wLoop_1_2 k_X k_I db) (combine (toW db) perms)). - - Definition Initialize_1_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_1_2 db k_S k_I k_Z ) (toW db); - [t_entries, perms] <-2 split wLoopRes; - xSet <- XSetSetup_1_2 k_X k_I db perms; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem G1_2_equiv : - Pr[G_gen Initialize_1_1] == Pr[G_gen Initialize_1_2]. - - unfold G_gen, Initialize_1_1, Initialize_1_2. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - Theorem Init_wLoop_1_2_spec : - forall d x x1 x2 x0 b0, - comp_spec - (fun x y => - fst (fst x) = fst y /\ snd x = snd y) - (Initialize_wLoop_1_1 d x x1 x2 x0 b0) - (Initialize_wLoop_1_2 d x x1 x2 b0). - - intuition. - unfold Initialize_wLoop_1_1. - unfold Initialize_wLoop_1_2. - comp_simp. - comp_skip. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - - Theorem Init_indLoop_1_2_spec : - forall x1 x2 x3 x0 b0 e, - comp_spec (fun p1 p2 => fst p1 = p2) - (OXT_EDBSetup_indLoop x1 x2 x3 x0 b0 e) - (OXT_EDBSetup_indLoop_1_2 x1 x2 x3 b0 e). - - intuition. - unfold OXT_EDBSetup_indLoop, OXT_EDBSetup_indLoop_1_2. - comp_simp. - comp_skip; try eapply (oneVector (lambda + lambda)). - eapply comp_spec_ret; intuition. - - Qed. - - eapply Init_indLoop_1_2_spec. - intuition. - eapply comp_spec_ret; intuition. - simpl. - eapply fst_split_eq_list_pred . - eauto. - - Qed. - - eapply Init_wLoop_1_2_spec. - intuition. - - remember (split a1) as z. - destruct z. - remember (split l0) as z. - destruct z. - remember (split b0) as z. - destruct z. - assert (fst (split (fst (split a1))) = fst (split b0)). - - eapply fst_split_eq_list_pred. - - Theorem list_pred_combine_l: - forall (A B : Set) (lsa : list A) (lsb : list B) (P : A -> B -> Prop), - list_pred P lsa lsb -> - forall (C : Set) (lsc : list C) (P1 : A -> C -> Prop) (P2 : B -> C -> Prop), - list_pred P1 lsa lsc -> - list_pred P2 lsb lsc -> - list_pred - (fun (p : A * B) (c : C) => - P (fst p) (snd p) /\ P1 (fst p) c /\ P2 (snd p) c) - (combine lsa lsb) lsc. - - induction 1; intuition. - inversion H; clear H; subst. - simpl. - econstructor. - - inversion H1; clear H1; subst. - inversion H2; clear H2; subst. - simpl. - econstructor. - intuition. - eauto. - - Qed. - - Show. - - Theorem list_pred_fst_split_l : - forall (A B C : Set) P (ls1 : list (A * B))(ls2 : list C), - list_pred (fun a b => P (fst a) b) ls1 ls2 -> - list_pred P (fst (split ls1)) ls2. - - induction 1; intuition; simpl in *. - econstructor. - destruct a1. - remember (split ls1) as z. - destruct z. - simpl in *. - econstructor; eauto. - Qed. - - eapply list_pred_fst_split_l. - - eapply list_pred_symm. - eapply list_pred_fst_split_l. - eapply list_pred_symm. - eapply list_pred_impl. - eauto. - intuition. - - rewrite <- Heqz in H7. - simpl in *. - rewrite <- Heqz0 in H7. - rewrite <- Heqz1 in H7. - simpl in *. - subst. - - comp_skip. - assert (flatten l3 = XSetSetup_1_2 x0 x1 d l5). - unfold XSetSetup_1_2. - f_equal. - eapply (@compMap_support _ _ (fun z1 z2 => in H4. - - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - - - Theorem list_pred_fst_split_r : - forall (A B C : Set) P (ls1 : list (A * B))(ls2 : list C), - list_pred (fun a b => P (fst a) b) ls1 ls2 -> - list_pred P (fst (split ls1)) ls2. - - induction 1; intuition; simpl in *. - econstructor. - destruct a1. - remember (split ls1) as z. - destruct z. - simpl in *. - econstructor; eauto. - Qed. - - SearchAbout list_pred zip. - - Qed. - - - - - (* Step 1.1: switch to the G2 initialization loop and split the XSet setup and TSet setup routines *) - Definition Initialize_1_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, perms] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db perms; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Theorem Init_indLoop_2_spec : - forall (x0 x1 x2 x a0 : Bvector lambda) b b1, - comp_spec - (fun (x : Bvector (lambda + lambda) * nat) (y : Bvector (lambda + lambda)) => - fst x = y -(* /\ snd x = g ^ (F_P x0 b * F_P x1 (Vector.to_list a0)) *) ) - (OXT_EDBSetup_indLoop x1 x2 (F x b) x0 b (a0, b1)) - (Initialize_indLoop_2 x1 x2 (F x b) b (a0, b1)). - - intuition. - unfold OXT_EDBSetup_indLoop, Initialize_indLoop_2. - comp_simp. - comp_skip; try eapply (oneVector). - eapply comp_spec_ret; intuition. - - Qed. - - - - Theorem Init_wLoop_2_spec : - forall d x x1 x2 b x0, - comp_spec - (fun (x : list (Vector.t bool (lambda + lambda)) * list nat) y => - fst x = fst y) - (OXT_EDBSetup_wLoop d x x1 x2 x0 b) - (Initialize_wLoop_2 d x x1 x2 b). - - intuition. - unfold OXT_EDBSetup_wLoop, Initialize_wLoop_2. - comp_simp. - unfold shuffle. - inline_first. - comp_skip. - rewrite permute_length_eq. - - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - - eapply comp_spec_seq; try eapply (nil, nil). - rewrite (@RndPerm_In_support_length (length (lookupInds d b))); intuition. - - eapply (@compMap_spec _ _ _ _ _ _ _ (fun (x : Bvector (lambda + lambda) * nat) (y : Bvector (lambda + lambda)) => - fst x = y)). - eapply list_pred_eq. - intuition. - subst. - eapply Init_indLoop_2_spec. - - intuition. - eapply comp_spec_ret; intuition. - - eapply fst_split_eq_list_pred. - trivial. - - Grab Existential Variables. - apply nat. - - Qed. - - Theorem TSet_inhabited : TSet. - - assert (TSet * TSetKey). - eapply comp_base_exists. - eapply TSetSetup. - eapply nil. - intuition. - Qed. - - Hint Resolve TSet_inhabited : inhabited. - - Require Import Permutation. - - Theorem G1_1_equiv : - Pr[G_gen Initialize_1] == Pr[G_gen Initialize_1_1]. - - unfold G_gen, Initialize_1, Initialize_1_1. - do 6 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; - eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - eapply Init_wLoop_2_spec. - intuition. - remember (split a1) as z. - destruct z. - - assert (fst (split a1) = b0). - eapply fst_split_eq_list_pred. - eauto. - rewrite <- Heqz in H7. - simpl in H7. - subst. - - comp_skip. - - eapply (@compMap_support _ _ (fun x y => Permutation (snd y) (XSetSetup_wLoop_2 x0 x1 d x))) in H4. - - assert (Permutation (flatten l1) (XSetSetup_2 x0 x1 d)). - unfold XSetSetup_2. - - Theorem Permutation_flatten_list_pred : - forall (A : Set)(ls1 ls2 : list (list A)), - list_pred (fun a b => Permutation a b) ls1 ls2 -> - Permutation (flatten ls1) (flatten ls2). - - induction 1; intuition; simpl in *. - eapply Permutation_app'; eauto. - Qed. - - eapply Permutation_flatten_list_pred. - - - - specialize (@list_pred_snd_split_irr _ (list (Bvector (lambda + lambda))) _ (fun x y => Permutation y (XSetSetup_wLoop_2 x0 x1 d x))); intuition. - - eapply list_pred_impl. - eapply list_pred_map_r. - eapply list_pred_symm. - eapply H9. - eapply H4. - - Theorem pair_snd_gen : - forall (A B : Type)(p : A * B) a b, - (a, b) = p -> - b = (snd p). - - intuition. - pairInv. - trivial. - - Qed. - - eapply pair_snd_gen. - eauto. - intuition. - destruct H12. - intuition. - subst. - trivial. - - - - SearchAbout snd. - - Check split. - Theorem split_eq_gen : - forall (A B : Type)(l1 l2 : list (A * B)), - l1 = l2 -> - split l1 = split l2. - - intuition. subst. intuition. - - Qed. - - erewrite split_eq_gen. - rewrite <- Heqz. - - remember ((@split - (list (Bvector (plus (posnatToNat lambda) (posnatToNat lambda)))) - (list nat) a1)) as z. - remember (split a1) as z. - rewrite <- Heqz. - - - SearchAbout Permutation flatten. - - f_equal. - assert (l1 = snd (split a1)). - rewrite <- Heqz. - trivial. - subst. - - - eapply snd_split_eq_map_list_pred. - eapply list_pred_symm. - eauto. - rewrite H9. - clear H9. - eapply comp_spec_eq_refl. - - intuition. - - unfold OXT_EDBSetup_wLoop in *. - repeat simp_in_support. - - eapply (@compMap_support _ _ - (fun (z1 : Identifier lambda * nat) (z2 : Bvector (lambda + lambda) * nat) - => snd z2 = g^((F_P x0 a2) * (F_P x1 (Vector.to_list (fst z1)))))) in H14. - - eapply snd_split_eq_map_list_pred. - eapply list_pred_symm. - - - specialize (@compMap_support _ _ - (fun z y => @snd (Blist) nat y = - XSetSetup_indLoop_2 x0 x1 (fst z) x) - _ - (fun z => OXT_EDBSetup_indLoop x1 x2 (F x a2) x0 a2 z)); intuition. - - Qed. - - (* Step 1.2 : factor out the tag calculation from the search loop *) - Definition GenTrans_1_2 (edb : EDB) (k : Key) (e : Query * Tag) : (list (Identifier lambda) * SearchTranscript) := - [q, stag] <-2 e; - t <- OXT_Search_ServerInit edb stag; - xscript <- map (OXT_Search_Loop edb k q) (combine (allNatsLt (length t)) t); - es <- getSomes (snd (split xscript)); - inds <- OXT_Search_ClientFinalize k q es; - (inds, (stag, xscript)). - - Definition Initialize_1_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - t_entries <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - xSet <- XSetSetup_2 k_X k_I db; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_1_2 edb key) (combine q stags); - ret (edb, searchRes). - - Theorem G1_2_equiv : - Pr[G_gen Initialize_1_1] == Pr[G_gen Initialize_1_2]. - - Admitted. - - (* simplify and re-order the transcript generation procedure to get to game 2*) - - Theorem G2_equiv : - Pr[G_gen Initialize_1_2] == Pr[G2]. - - Admitted. - - - (* Step 3: prepare to replace all the PRFs with random functions. *) - (* We use a composite oracle to keep the complexity under control. *) - - Notation "'query' v" := (OC_Query _ v) (at level 79) : comp_scope. - - Instance sum_EqDec : - forall (A B : Set)(eqda : EqDec A)(eqdb : EqDec B), - EqDec (A + B). - - intuition. - exists - (fun x y => - match x with - | inl a => - match y with - | inl a' => eqb a a' - | inr _ => false - end - | inr b => - match y with - | inl _ => false - | inr b' => eqb b b' - end - end). - - intuition; try discriminate. - f_equal. - eapply eqb_leibniz. - trivial. - inversion H; clear H; subst. - eapply eqb_leibniz. - trivial. - - f_equal. - eapply eqb_leibniz. - trivial. - inversion H; clear H; subst. - eapply eqb_leibniz. - trivial. - Qed. - - Definition mkCompositeOracle - {A1 A2 B1 B2 S1 S2 : Set} - {eqdb1 : EqDec B1}{eqdb2 : EqDec B2} - {eqds1 : EqDec S1}{eqds2 : EqDec S2} - (o1 : S1 -> A1 -> Comp (B1 * S1)) - (o2 : S2 -> A2 -> Comp (B2 * S2)) - (s : S1 * S2)(z : A1 + A2) := - match z with - | inl a => [b, s'] <-$2 o1 (fst s) a; ret (inl b, (s', snd s)) - | inr a => [b, s'] <-$2 o2 (snd s) a; ret (inr b, (fst s, s')) - end. - - Definition toL(A B : Type)(def : A)(x : A + B) : A := - match x with - | inl a => a - | inr _ => def - end. - - Definition toR(A B : Type)(def : B)(x : A + B) : B := - match x with - | inl _ => def - | inr b => b - end. - - (* - / \ - / \ / \ - S X I Z *) - - Definition D_O := (Blist + Blist + (Blist + Blist))%type. - Definition R_O := (Bvector lambda + nat + (nat + nat))%type. - - Definition QS (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QX (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - Definition QI (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QZ (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - - Definition toS(r : R_O) := - x <- toL (inl nat (oneVector lambda)) r; - toL (oneVector lambda) x. - - Definition toX(r : R_O) := - x <- toL (inl nat (oneVector lambda)) r; - toR O x. - - Definition toI(r : R_O) := - x <- toR (inl nat O) r; - toL O x. - - Definition toZ(r : R_O) := - x <- toR (inl nat O) r; - toR O x. - - Definition sQuery(x : Blist) := - qRes <--$ query (QS x); - $ ret (toS qRes). - - Definition xQuery(x : Blist) := - qRes <--$ query (QX x); - $ ret (toX qRes). - - Definition iQuery(x : Blist) := - qRes <--$ query (QI x); - $ ret (toI qRes). - - Definition zQuery(x : Blist) := - qRes <--$ query (QZ x); - $ ret (toZ qRes). - - Definition Initialize_indLoop_3 k_E w (e : Identifier lambda * nat) : OracleComp D_O R_O (Bvector (lambda + lambda)) := - [ind, c] <-2 e; - e <--$$ Enc k_E ind; - xind <--$ xQuery (Vector.to_list ind); - z <--$ zQuery (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y)). - - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - Definition Initialize_wLoop_3 db w : OracleComp D_O R_O (list (Bvector (lambda + lambda))):= - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$ sQuery w; - oc_compMap _ (Initialize_indLoop_3 k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))). - - Definition XSetSetup_indLoop_3 w (ind : Bvector lambda) := - e <--$ xQuery w; - xind <--$ iQuery (Vector.to_list ind); - $ ret (g^(e * xind)). - - Definition XSetSetup_wLoop_3 db w := - oc_compMap _ (XSetSetup_indLoop_3 w) (lookupInds db w). - - Definition XSetSetup_3 db := - ls <--$ oc_compMap _ (XSetSetup_wLoop_3 db) (toW db); - $ ret removeDups _ (flatten ls). - - Definition GenTrans_3 (edb : EDB) (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <--$ xQuery x; - xtokens <--$ oc_compMap _(fun (c : nat) => exp <--$ zQuery (s ++ c); $ ret g^(e * exp)) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <--$ sQuery s; - inds <- map (Dec k_E) es; - $ ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_3_oc (db : DB)(q : list Query) := - t_entries <--$ oc_compMap _ (Initialize_wLoop_3 db) (toW db); - xSet <--$ XSetSetup_3 db; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - edb <- (tSet, xSet); - stags <--$$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_3 edb) (combine q stags); - $ ret (edb, searchRes). - - Definition Initialize_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o1 := mkCompositeOracle (f_oracle F _ k_S) (f_oracle F_P _ k_X) in - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_3_oc db q) _ _ o (tt, tt, (tt, tt)); - ret r. - - Definition G3 := G_gen Initialize_3. - - Theorem G3_equiv : - Pr[G2] == Pr[G3]. - - Admitted. - - Check F_P. - - (* Step 4: replace all PRFs with random functions. *) - Definition Initialize_4 (db : DB)(q : list Query) := - let o1 := mkCompositeOracle (@randomFunc Keyword _ (Rnd lambda) _ ) (@randomFunc Blist _ (RndNat order) _) in - let o2 := mkCompositeOracle (@randomFunc Blist _ (RndNat order) _) (@randomFunc Blist _ (RndNat order) _) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_3_oc db q) _ _ o (nil, nil, (nil, nil)); - ret r. - - Definition G4 := G_gen Initialize_4. - - Theorem G4_equiv : - Pr[G3] == Pr[G4]. - - Admitted. - - (* Step 5: replace k_E with a randomly-selected value. *) - Definition Initialize_wLoop_5 db w : OracleComp D_O R_O (list (Bvector (lambda + lambda))):= - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$$ {0, 1}^lambda; - oc_compMap _ (Initialize_indLoop_3 k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))). - Definition Initialize_5_oc (db : DB)(q : list Query) := - t_entries <--$ oc_compMap _ (Initialize_wLoop_5 db) (toW db); - xSet <--$ XSetSetup_3 db; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - edb <- (tSet, xSet); - stags <--$$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_3 edb) (combine q stags); - $ ret (edb, searchRes). - - Definition Initialize_5 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - let o1 := mkCompositeOracle (f_oracle F _ k_S) (f_oracle F_P _ k_X) in - let o2 := mkCompositeOracle (f_oracle F_P _ k_I) (f_oracle F_P _ k_Z) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_5_oc db q) _ _ o (tt, tt, (tt, tt)); - ret r. - - Definition G5 := G_gen Initialize_5. - - Theorem G5_equiv : - Pr[G4] == Pr[G5]. - - Admitted. - - - (* Step 6: replace the answer computation with direct answer lookups. *) - Definition GenTrans_6 (db : DB)(edb : EDB) (e : Query * Tag) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <--$ xQuery x; - xtokens <--$ oc_compMap _(fun (c : nat) => exp <--$ zQuery (s ++ c); $ ret g^(e * exp)) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - inds <- intersect (EqDec_dec _) (lookupInds db s) (lookupInds db x); - $ ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_6_oc (db : DB)(q : list Query) := - t_entries <--$ oc_compMap _ (Initialize_wLoop_5 db) (toW db); - xSet <--$ XSetSetup_3 db; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - edb <- (tSet, xSet); - stags <--$$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_6 db edb) (combine q stags); - $ ret (edb, searchRes). - - Definition Initialize_6 (db : DB)(q : list Query) := - let o1 := mkCompositeOracle (@randomFunc Keyword _ (Rnd lambda) _ ) (@randomFunc Blist _ (RndNat order) _) in - let o2 := mkCompositeOracle (@randomFunc Blist _ (RndNat order) _) (@randomFunc Blist _ (RndNat order) _) in - let o := mkCompositeOracle o1 o2 in - [r, _] <-$2 (Initialize_6_oc db q) _ _ o (nil, nil, (nil, nil)); - ret r. - - Definition G6 := G_gen Initialize_6. - - Theorem G6_close : - | Pr[G5] - Pr[G6] | <= 0. - - (* not really 0, but something small. *) - (* These games are only different if we end up with collisions in some of the random functions or other issues (see below) *) - - Admitted. - - (* In the argument above, we mostly need to account for things that can go wrong, but we (may) also need to permute the initialization so the results come back in the correct order. We can correct the order by choosing the identity permutation. - -Things that can go wrong: - -1) TSet is incorrect -2) token collision: g ^ (F1(x) * F2(s ++ c)) for random functions F1 and F2. We can show that this is a random function of (x, s, c), and therefore collisions are unlikely. -3) XSet collision: there are two (keyword, index) pairs, (w1, i1) (w2, i2) for which g^(F1(w1) * F2(i1)) = g^(F1(w2) * F2(i2)). Like before, use an argument that treats this whole thing as a random function -4) Decrypt function is incorrect (probability 0) - *) - - (* G6 above is exacly G1 in the paper*) - - - Definition whydoihavetodothingslikethistofixmyspacing := tt. - - (* Step 6: since - - - (* Combine 2 (or more) oracles into a single "composite oracle". We get back: - 1) The oracle itself - 2) A family of functions that injects arguments into the composite argument type. - 3) A function that takes a composite response and a composite argument and returns the desired response - *) - (* Inject argument into sum *) - (* Give sum to composite oracle, get back product *) - - Inductive BinTreeStruct : Set := - | btsLeaf : BinTreeStruct - | btsNode : BinTreeStruct -> BinTreeStruct -> BinTreeStruct. - - - Inductive BinTree(A : Type) : BinTreeStruct -> Type := - | btLeaf : A -> BinTree A btsLeaf - | btNode : - forall t1 t2, - BinTree A t1 -> - BinTree A t2 -> - BinTree A (btsNode t1 t2). - - Fixpoint BinTree_interp(A : Type)(op : A -> A -> A)(b : BinTreeStruct)(t : BinTree A b) : A := - match t with - | btLeaf a => a - | btNode _ _ t1 t2 => op (BinTree_interp op t1) (BinTree_interp op t2) - end. - - Definition BinTree_sum := BinTree_interp sum. - Definition BinTree_prod := BinTree_interp prod. - - Inductive BinTreePath : BinTreeStruct -> Type := - | btpThis : BinTreePath btsLeaf - | btpLeft: - forall t1 t2 (p : BinTreePath t1), - BinTreePath (btsNode t1 t2) - | btpRight: - forall t1 t2 (p : BinTreePath t2), - BinTreePath (btsNode t1 t2). - - Definition binTreeLeft(bt :BinTreeStruct) : BinTreeStruct := - match bt with - | btsLeaf => btsLeaf - | btsNode t1 t2 => t1 - end. - - Definition binTreeRight(bt :BinTreeStruct) : BinTreeStruct := - match bt with - | btsLeaf => btsLeaf - | btsNode t1 t2 => t2 - end. - - Fixpoint treeEntry(T : BinTreeStruct)(A : Type)(X : BinTree A T) : forall (p : BinTreePath T), A := - match X with - | btLeaf x => fun p => x - | btNode t1 t2 st1 st2 => fun p => - (match p in (BinTreePath T') return ( - (BinTreePath (binTreeLeft T') -> A) -> - (BinTreePath (binTreeRight T') -> A) -> - match T' return Type with - | btsLeaf => unit - | _ => A - end - ) with - | btpThis => fun f1 f2 => tt - | btpLeft t1' t2' p' => fun f1 f2 => - f1 p' - | btpRight t1' t2' p' => fun f1 f2 => - f2 p' - end) (treeEntry st1) (treeEntry st2) - end. - - (* Next: make a type for heteregeneous trees like in CPDT *) - - Fixpoint sumToTreePath(T : BinTree)(A : SetTree T) : forall (x : SetTree_sum A), BinTreePath T := - match A in (SetTree T') return (forall (x : SetTree_sum A), BinTreePath T') with - | stLeaf _ => fun x => btpThis - | stNode t1 t2 st1 st2 => - fun x => - match x with - | inl x' => btpLeft t2 (sumToTreePath st1 x') - | inr x' => btpRight t1 (sumToTreePath st2 x') - end - end. - - Fixpoint treePathToProdEntry(T : BinTree) : - forall (A : SetTree T)(p : BinTreePath T)(x : SetTree_prod A), treeEntryType A p := - match T with - | btLeaf => - fun A => match A with - | - | _ => tt - end. - - Fixpoint treePathToProdEntry(T : BinTree)(A : SetTree T) : - forall (p : BinTreePath T)(x : SetTree_prod A), treeEntryType A p := - match A in (SetTree T') return - (forall (p : BinTreePath T')(x : SetTree_prod A), treeEntryType A p) with - | stLeaf _ => fun p x => x - | stNode t1 t2 st1 st2 => - fun p x => - match p in (BinTreePath T') return - ((forall (p': BinTreePath (binTreeLeft T')), (SetTree_prod st1) -> treeEntryType st1 p') -> - (forall (p': BinTreePath t1), (SetTree_prod st1) -> treeEntryType st1 p') -> - unit) - with - | btpThis => fun f1 f2 => tt - | btpLeft t1' t2' p' => - fun f1 f2 => - f1 p' (fst x) - | btpRight t1' t2' p' => fun f1 f2 => - (treePathToProdEntry st2) p' (snd x) - end - end. - - - (match p in (BinTreePath T') return ( - (BinTreePath (binTreeLeft T') -> Set) -> - (BinTreePath (binTreeRight T') -> Set) -> Set) with - | btpThis => fun f1 f2 => unit - | btpLeft t1' t2' p' => fun f1 f2 => - f1 p' - | btpRight t1' t2' p' => fun f1 f2 => - f2 p' - end) (treeEntryType st1) (treeEntryType st2) - end. - - - Fixpoint toProdEntry(T : BinTree) : - forall (A B : SetTree T)(a : SetTree_sum A)(b : SetTree_prod B), := - match T with - | btLeaf => tt - | btNode t1 t2 => tt - end. - - Definition COracle(S A B : SetTree) := - SetTree_prod S -> SetTree_sum A -> Comp (SetTree_prod B * SetTree_prod S). - - Definition mkCompositeOracle - {A1 A2 B1 B2 S1 S2 : Set} - {eqdb1 : EqDec B1}{eqdb2 : EqDec B2} - {eqds1 : EqDec S1}{eqds2 : EqDec S2} - (d1 : B1)(d2 : B2) - (o1 : S1 -> A1 -> Comp (B1 * S1)) - (o2 : S2 -> A2 -> Comp (B2 * S2)) - (s : S1 * S2)(z : A1 + A2) := - match z with - | inl a => [b, s'] <-$2 o1 (fst s) a; ret ((b, d2), (s', snd s)) - | inr a => [b, s'] <-$2 o2 (snd s) a; ret ((d1, b), (fst s, s')) - end. - - - - Definition OC_CQuery (A B : SetTree) : - forall (z : SetTree_sum A), - OracleComp (SetTree_sum A) (SetTree_prod B) - (match z return Set with - | inl _ => B1 - | inr _ => B2 - end) := - fun z => - match z with - | inl a => [b, _] <--$2 query z; $ ret b - | inr a => [_, b] <--$2 query z; $ ret b - end. - - (* - / \ - / \ / \ - S X I Z *) - - Definition QS (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QX (x : Blist) := @inl (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - Definition QI (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inl Blist Blist x). - Definition QZ (x : Blist) := @inr (Blist + Blist) (Blist + Blist) (@inr Blist Blist x). - - Definition D_O := (Blist + Blist + (Blist + Blist))%type. - Definition R_O := (Bvector lambda * nat * (nat * nat))%type. - Definition to_rVal (d : D_O)(r : R_O) := - match d with - | inl x => - match x with - | inl _ => fst (fst r) - | inr _ => snd (fst r) - end - | inr x => - match x with - | inl _ => fst (snd r) - | inr _ => snd (snd r) - end - end. - - - - - - - - - - Notation "'cquery' v" := (OC_CQuery v) (at level 79) : comp_scope. - - Definition Initialize_indLoop_3 k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E ind; - xind <--$ cquery (QI (Vector.to_list ind)); - z <--$ cquery (QZ (w ++ c)); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y)). - - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - - Definition Initialize_wLoop_3 db w : OracleComp D_O R_O (list (Bvector (lambda + lambda))):= - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$ cquery (QS w); - oc_compMap _ (Initialize_indLoop_3 k_E w) (combine (permute (oneVector lambda) inds sigma) (allNatsLt (length inds))). - - Definition XSetSetup_indLoop_2 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - g^(e * xind). - - Definition XSetSetup_wLoop_2 k_X k_I db w := - map (XSetSetup_indLoop_2 k_X k_I w) (lookupInds db w). - - Definition XSetSetup_2 k_X k_I db := - removeDups _ (flatten (map (XSetSetup_wLoop_2 k_X k_I db) (toW db))). - - - Definition ServerSearchLoop_2 (edb : EDB) e := - [xtoken, t_c] <-2 e; - [_, xSet] <-2 edb; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition ServerSearch_2 edb (xtokens : list nat) t := - map (ServerSearchLoop_2 edb) (combine xtokens t). - - Definition GenTrans_2 (edb : EDB) k_X k_Z k_S (e : Query * Tag) : (list (Identifier lambda) * SearchTranscript) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 edb xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - t_entries <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - xSet <- XSetSetup_2 k_X k_I db; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_2 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - - Theorem G2_equiv : - Pr[G_gen Initialize_1_3] == Pr[G2]. - - unfold G2, G_gen, Initialize_1_3, Initialize_2. - do 9 (comp_skip; comp_simp). - eapply evalDist_ret_eq. - f_equal. - eapply map_ext. - intuition. - unfold GenTrans_1_3, GenTrans_2. - comp_simp. - eapply map_eq. - - Qed. - - - Definition OXT_Search_Loop_Client (k : Key)(q : Query) c := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [s, x] <-2 q; - computeXToken c k_Z k_X s x. - - Definition OXT_Search_Loop_Server (edb : EDB) xtoken t_c := - [_, xSet] <-2 edb; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition OXT_Search_Loop edb k q (e : nat * Bvector (lambda + lambda)) := - [c, t_c] <-2 e; - xtoken <- OXT_Search_Loop_Client k q c ; - answer <- OXT_Search_Loop_Server edb xtoken t_c; - (xtoken, answer). - - Definition Initialize_1_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - t_entries <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - xSet <- XSetSetup_2 k_X k_I db; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_1_3 edb key) (combine q stags); - ret (edb, searchRes). - - Theorem G1_3_equiv : - Pr[G_gen Initialize_1_2] == Pr[G_gen Initialize_1_3]. - - Admitted. - - - Theorem G0_equiv : - Pr[SSE_Sec_NA_Real OXT_EDBSetup _ OXT_Search A1 A2 ] == Pr[G0]. - - unfold SSE_Sec_NA_Real, G0. - comp_skip. - comp_simp. - unfold OXT_EDBSetup, Initialize_0. - do 4 (inline_first; comp_skip). - - inline_first. - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - - Theorem Init_indLoop_0_spec : - forall x1 x2 x b x0 a0 b1, - comp_spec - (fun a b => fst a = b) - (OXT_EDBSetup_indLoop x1 x2 (F x b) x0 b (a0, b1)) - (Initialize_indLoop_0 x1 x2 (F x b) b (a0, b1)). - - intuition. - unfold OXT_EDBSetup_indLoop, Initialize_indLoop_0. - comp_simp. - comp_skip; try eapply (oneVector). - eapply comp_spec_ret; intuition. - - Qed. - - Theorem Init_wLoop_0_spec : forall d x x1 x2 b x0, - comp_spec - (fun a b => fst a = b) - (OXT_EDBSetup_wLoop d x x1 x2 x0 b) - (Initialize_wLoop_0 d x x1 x2 b). - - intuition. - unfold OXT_EDBSetup_wLoop, Initialize_wLoop_0. - comp_simp. - unfold shuffle. - inline_first. - comp_skip. - rewrite permute_length_eq. - - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - - eapply comp_spec_seq; try eapply (nil, nil). - rewrite (@RndPerm_In_support_length (length (lookupInds d b))); intuition. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - eapply Init_indLoop_0_spec. - - intuition. - eapply comp_spec_ret; intuition. - - Theorem fst_split_eq_list_pred : - forall (A B : Set)(ls1 : list (A * B))(ls2 : list A), - list_pred (fun a b => fst a = b) ls1 ls2 -> - fst (split ls1) = ls2. - - induction 1; intuition; simpl in *. - subst. - destruct a1. - remember (split ls1) as x. - destruct x. - simpl in *. - f_equal. - Qed. - - eapply fst_split_eq_list_pred. - trivial. - - Grab Existential Variables. - (* Not sure where this came from. *) - apply nat. - - Qed. - - eapply Init_wLoop_0_spec. - intuition. - - remember (split a0) as z. - comp_simp. - inline_first. - assert (l0 = b). - admit. - subst. - comp_skip. - inline_first. - - Theorem OXT_Seqrch_0_equiv : - forall a x4 x0 x1 x2 b0 a2 l2, - comp_spec - (fun (b : Query * Tag) (c : list (Identifier lambda) * SearchTranscript) => - OXT_Search_0 (a2, l2) (x4, x0, x1, x2, b0) b = c) - (z <-$ TSetGetTag b0 (fst a); ret (a, z)) - (OXT_Search (a2, l2) (x4, x0, x1, x2, b0) a). - - intuition. - unfold OXT_Search. - unfold OXT_Search_ClientInit. - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - eapply comp_spec_ret. - unfold OXT_Search_0. - comp_simp. - reflexivity. - Qed. - - Theorem OXT_Search_0_loop_equiv : - forall (l : list Query) a2 l2 x4 x0 x1 x2 b0, - comp_spec - eq - (compMap _ (OXT_Search (a2, l2) (x4, x0, x1, x2, b0)) l) - (ls' <-$ compMap _ (fun x => stag <-$ TSetGetTag b0 (fst x); ret (x, stag)) l; - ret (map (fun x => OXT_Search_0 (a2, l2) (x4, x0, x1, x2, b0) x) ls')). - - intuition. - - eapply comp_spec_eq_symm. - eapply comp_spec_eq_trans. - specialize (@compMap_map_fission_eq Query - (Query * Tag) - (list (Identifier lambda) * SearchTranscript) - (list (Identifier lambda) * SearchTranscript) - _ _ _ - l - (fun x => z <-$ TSetGetTag b0 (fst x); ret (x, z)) - (OXT_Search (a2, l2) (x4, x0, x1, x2, b0)) - (OXT_Search_0 (a2, l2) (x4, x0, x1, x2, b0)) - (fun x => x)); intuition. - eapply H. - - intuition. - eapply OXT_Seqrch_0_equiv. - eapply comp_spec_consequence. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - eapply comp_spec_right_ident. - intuition. - eapply list_pred_eq_impl_eq. - trivial. - Qed. - - Show. - - eapply comp_spec_eq_trans. - eapply comp_spec_seq_eq; eauto with inhabited. - eapply OXT_Search_0_loop_equiv. - intuition. - eapply comp_spec_eq_refl. - inline_first. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - assert (list_pred (fun a b => fst a = b) l (fst (split l))). - admit. - eauto. - intuition. - subst. - simpl. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - eapply comp_base_exists; eauto. - eapply comp_base_exists; eauto. - assert (comp_spec (fun a b => snd a = b) (ret (a3, b1, b2)) (ret b2)). - eapply comp_spec_ret; intuition. - eauto. - intuition. - comp_simp. - - Theorem comp_spec_eq_refl_gen : - forall (A : Set)(eqd : EqDec A)(c1 c2 : Comp A), - c1 = c2 -> - comp_spec eq c1 c2. - - intuition. - subst. - eapply comp_spec_eq_refl. - - Qed. - - eapply comp_spec_eq_refl_gen. - f_equal. - f_equal. - admit. - - f_equal. - f_equal. - unfold OXT_Search_0, GenTrans_0. - - unfold XSetSetup_0, XSetSetup_wLoop_0, XSetSetup_indLoop_0. - - Check compMap_support. - - Theorem compMap_support' - : forall (A B : Set) (eqd : EqDec B) - (c : A -> Comp B) (lsa : list A) (lsb : list B), - In lsb (getSupport (compMap _ c lsa)) -> - list_pred (fun a b => In b (getSupport (c a))) lsa lsb. - - - eapply compMap_support in H4. - - - unfold OXT_Search_0, GenTrans_0. - comp_simp. - - eapply comp_spec_ret. - - eapply comp_spec_ret; intuition. - assert ((fun a b => snd a = b) (a3, b1, b2) b2). - trivial. - eapply H15. - eapply comp_spec_eq_trans. - eapply comp_spec_eq_symm. - eapply compMap_fission_eq. - Check compMap_map_fission_eq. - - - Print OXT_Search. - - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_spec. - assert (list_pred (fun a b => fst a = b) l (fst (split l))). - admit. - eauto. - intuition. - subst. - - Qed. - - - Definition x := tt. - - - - - - (* Step 2: Replace PRF F with a random function. We can use a non-adaptively-secure PRF, since k_S isn't used after initialization. *) - - (* Step 3: inputs to the random function are unique---replace outputs with random values. *) - - (* Step 4-6: Replace F_P with random functions (one for each key). Use adapatively secure PRF. *) - - (* Step 7: Replace the "client finalize" step in the search protocol with a function that looks up the correct answers. Use the correctness of the encryption and T-Set schemes to show that this is close. *) - - (* Step 8: Use IND-CPA to replace encryptions of indices with encryptions of 0^lambda. This is possible because decryption never happens. *) - - (* Step 9: - - -End ESPADA_SSE_OXT. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_SSE_OXT_Games.v b/src/ESPADA/ESPADA_SSE_OXT_Games.v deleted file mode 100644 index 39472e6..0000000 --- a/src/ESPADA/ESPADA_SSE_OXT_Games.v +++ /dev/null @@ -1,876 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Set Implicit Arguments. - -Require Import Coq.ZArith.Zdigits. - -Require Import FCF. -Require Import CompFold. -Require Import GroupTheory. -Require Import RndPerm. -Require Import PRF. - -Require Import SSE. -Require Import TSet. - -Local Open Scope list_scope. -Local Open Scope type_scope. - -Section ESPADA_SSE_OXT_Games. - - Variable lambda : posnat. - Definition DB := DB lambda. - - Variable p : posnat. - Definition multModP (a b : nat) := - modNat (a * b) p. - - Context `{FCG : FiniteCyclicGroup nat}. - - Hypothesis op_multModP : - groupOp = multModP. - - (* PRFs *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - Variable F_P : Bvector lambda -> Blist -> nat. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - Hypothesis Enc_correct : - forall k p c, - In c (getSupport (Enc k p)) -> - Dec k c = p. - - - Instance nz_posnat_plus : - forall (x y : posnat), - nz (x + y). - - intuition. - unfold posnatToNat, natToPosnat. - destruct x. - destruct y. - econstructor. - omega. - - Qed. - - (* TSet *) - Variable TSet : Set. - Hypothesis TSet_EqDec : EqDec TSet. - Variable TSetKey : Set. - Hypothesis TSetKey_EqDec : EqDec TSetKey. - Variable Tag : Set. - Hypothesis Tag_EqDec : EqDec Tag. - Variable TSetSetup : T (pos (lambda + lambda)) -> Comp (TSet * TSetKey). - Variable TSetGetTag : TSetKey -> Keyword -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector (lambda + lambda))). - Variable Leakage_T : Set. - Hypothesis Leakage_T_EqDec : EqDec Leakage_T. - Variable L_T : T (pos (lambda + lambda)) -> list Keyword -> Leakage_T. - - Definition XSet := list nat. - Definition EDB := TSet * XSet. - Definition Key := Bvector lambda * Bvector lambda * Bvector lambda * Bvector lambda * TSetKey. - Definition Query := Keyword * Keyword. - - Definition natToBvector(n : nat) : Bvector lambda := - Z_to_binary lambda (Z.of_nat n). - - Definition natToBlist (n : nat) := - Vector.to_list (natToBvector n). - - Definition bvectorToNat(v : Bvector lambda) : nat := - Z.to_nat (binary_value _ v). - - Coercion natToBlist : nat >-> list. - - Definition lookupInds(db : DB)(w : Keyword) := - fold_left (fun ls p => if (in_dec (EqDec_dec _) w (snd p)) then (fst p :: ls) else ls) db nil. - - Local Open Scope group_scope. - - Definition OXT_EDBSetup_indLoop k_I k_Z k_E k_X w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - e <-$ Enc k_E ind; - xtag <- g^((F_P k_X w) * xind); - ret (Vector.append e (natToBvector y), xtag). - - Definition OXT_EDBSetup_wLoop db k_S k_I k_Z k_X w := - k_E <- F k_S w; - inds <- lookupInds db w; - inds <-$ shuffle _ inds; - indLoopRes <-$ compMap _ (OXT_EDBSetup_indLoop k_I k_Z k_E k_X w) (combine inds (allNatsLt (length inds))); - ret (split indLoopRes). - - Definition toW (db : DB) := - removeDups _ (flatten (snd (split db))). - - Definition OXT_EDBSetup (db : DB) : Comp (Key * (TSet * XSet)) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (OXT_EDBSetup_wLoop db k_S k_I k_Z k_X) (toW db); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - ret ((k_S, k_X, k_I, k_Z, k_T), (tSet, xSet)). - - Definition OXT_Search_ClientInit(k : Key)(q : Query) := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [w1, _] <-2 q; - TSetGetTag k_T w1. - - Definition OXT_Search_ServerInit(edb : EDB) stag := - [tSet, _] <-2 edb; - TSetRetrieve tSet stag. - - Definition computeXToken c k_Z k_X s x := - g^((F_P k_Z (s ++ c)) * (F_P k_X x)). - - Definition OXT_Search_Loop_Client (k : Key)(q : Query) c := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [s, x] <-2 q; - computeXToken c k_Z k_X s x. - - Fixpoint splitVector(A : Set)(n m : nat) : Vector.t A (n + m) -> (Vector.t A n * Vector.t A m) := - match n with - | 0 => - fun (v : Vector.t A (O + m)) => (@Vector.nil A, v) - | S n' => - fun (v : Vector.t A (S n' + m)) => - let (v1, v2) := splitVector _ _ (Vector.tl v) in - (Vector.cons _ (Vector.hd v) _ v1, v2) - end. - - Definition OXT_Search_Loop_Server (edb : EDB) xtoken t_c := - [_, xSet] <-2 edb; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition OXT_Search_Loop edb k q (e : nat * Bvector (lambda + lambda)) := - [c, t_c] <-2 e; - xtoken <- OXT_Search_Loop_Client k q c ; - answer <- OXT_Search_Loop_Server edb xtoken t_c; - (xtoken, answer). - - Definition OXT_Search_ClientFinalize (k : Key) (q : Query) es := - [k_S, k_X, k_I, k_Z, k_T] <-5 k; - [w1, ws] <-2 q; - k_E <- F k_S w1; - map (Dec k_E) es. - - Fixpoint getSomes(A : Type)(ls : list (option A)) := - match ls with - | nil => nil - | o :: ls' => - match o with - | None => (getSomes ls') - | Some a => a :: (getSomes ls') - end - end. - - - Definition SearchTranscript := (Tag * list (nat * option (Bvector lambda)))%type. - - Definition OXT_Search (edb : EDB) (k : Key) (q : Query) : Comp (list (Identifier lambda) * SearchTranscript) := - stag <-$ OXT_Search_ClientInit k q; - t <- OXT_Search_ServerInit edb stag; - xscript <- map (OXT_Search_Loop edb k q) (combine (allNatsLt (length t)) t); - es <- getSomes (snd (split xscript)); - inds <- OXT_Search_ClientFinalize k q es; - ret (inds, (stag, xscript)). - - Definition dbSize(db : DB) := - fold_left (fun n e => n + length (snd e))%nat db 0%nat. - - Require Import RndListElem. - Check firstIndexOf. - - Definition equalityPattern(ls : list Keyword) := - map (fun x => firstIndexOf (EqDec_dec _) ls x 0) ls. - - Definition sizePattern(db : DB)(ls : list Keyword) := - map (fun x => length (lookupInds db x)) ls. - - Definition lookupIndsConj(db : DB)(w : Query) := - [s, x] <-2 w; - inds <- lookupInds db s; - intersect (EqDec_dec _) inds (lookupInds db x). - - Definition resultsPattern(db : DB)(q: list Query) := - map (lookupIndsConj db) q. - - Definition condIntPattern(db : DB) (i j : nat) (qi qj : Query):= - r <- intersect (EqDec_dec _) (lookupInds db (fst qi)) (lookupInds db (fst qj)); - if (eq_nat_dec i j) then nil else - if (eqb (snd qi) (snd (qj))) then r else nil. - - Definition condIntPatternTable(db : DB)(q : list Query) := - map (fun e => [i,qi] <-2 e; - map - (fun e => [j,qj] <-2 e; condIntPattern db i j qi qj) - (combine (allNatsLt (length q)) q)) - (combine (allNatsLt (length q)) q). - - Definition L_OXT(db : DB)(q : list Query) := - n <- dbSize db; - s_bar <- equalityPattern (fst (split q)); - sP <- sizePattern db (fst (split q)); - rP <- resultsPattern db q; - iP <- condIntPatternTable db q; - (n, s_bar, sP, rP, iP). - - Require Import RndGrpElem. - - Definition L_cLoop (k : Bvector lambda)(c : nat) := - y <-$ RndG; - e <-$ Enc k (Vector.const false lambda); - ret (Vector.append (natToBvector y) e). - - Definition L_wLoop (db : DB)(w : Keyword) := - k <-$ {0, 1}^lambda; - t_entries <-$ compMap _ (L_cLoop k) (allNatsLt (length (lookupInds db w))); - ret (w, t_entries). - - Definition L (db : DB)(q : list Query) := - bigT <-$ compMap _ (L_wLoop db) (toW db); - s <- fst (unzip q); - tSetAnswers <- map (lookupAnswers (pos (lambda + lambda))%nat bigT) s; - ret (L_OXT db q, L_T bigT s, tSetAnswers ). - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - (* Step 1: simplify and put the game into the "Initialize" form of the paper. *) - Definition Initialize_1 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (OXT_EDBSetup_wLoop db k_S k_I k_Z k_X) (toW db); - [t_entries, xSet_raw] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - xSet <- flatten xSet_raw; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - key <- (k_S, k_X, k_I, k_Z, k_T); - searchRes <-$ compMap _ (OXT_Search edb key) q; - ret (edb, searchRes). - - Definition G_gen (init : DB -> list Query -> Comp (EDB * (list (list (Bvector lambda) * SearchTranscript)))) := - [db, q, s_A] <-$3 A1; - [edb, ls] <-$2 init db q; - A2 s_A edb (snd (split ls)). - - Definition G1 := G_gen Initialize_1. - - Theorem SSE_G1_equiv : - Pr[SSE_Sec_NA_Real OXT_EDBSetup _ OXT_Search A1 A2 ] == Pr[G1]. - Abort. - - - (* Step 2: simplify and split some loops. This gets us close to game G0 in the paper, except we still -look up answers in the real way. *) - - Definition Initialize_indLoop_2 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y)). - - Definition Initialize_wLoop_2 db k_S k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <- F k_S w; - indLoopRes <-$ compMap _ (Initialize_indLoop_2 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition XSetSetup_indLoop_2 k_X k_I w (ind : Bvector lambda) := - e <- F_P k_X w; - xind <- F_P k_I (Vector.to_list ind); - g^(e * xind). - - Definition XSetSetup_wLoop_2 k_X k_I db e := - [w, sigma] <-2 e; - map (XSetSetup_indLoop_2 k_X k_I w) (permute (lookupInds db w) sigma). - - Definition XSetSetup_2 k_X k_I db sigmas := - flatten (map (XSetSetup_wLoop_2 k_X k_I db) (combine (toW db) sigmas)). - - Definition ServerSearchLoop_2 xSet e := - [xtoken, t_c] <-2 e; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some e) else None. - - Definition ServerSearch_2 xSet (xtokens : list nat) t := - map (ServerSearchLoop_2 xSet) (combine xtokens t). - - Definition GenTrans_2 (edb : EDB) k_X k_Z k_S (e : Query * Tag) : (list (Identifier lambda) * SearchTranscript) := - [q, stag] <-2 e; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - t <- TSetRetrieve tSet stag; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_2 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags <-$ compMap _ (TSetGetTag k_T) (fst (split q)); - searchRes <- map (GenTrans_2 edb k_X k_Z k_S) (combine q stags); - ret (edb, searchRes). - - Definition G2 := G_gen Initialize_2. - - Theorem G1_G2_equiv : - Pr[G1] == Pr[G2]. - Abort. - - - (* Step 3 : replace the TSetRetrieve and use the actual values instead. We make this transformation by applying the TSet correctness definition. *) - Definition GenTrans_3 (edb : EDB) k_X k_Z k_S (e : Query * (Tag * list (Bvector (lambda + lambda)))) : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_2 xSet xtokens t; - es <- getSomes res; - k_E <- F k_S s; - inds <- map (Dec k_E) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_3 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_3 edb k_X k_Z k_S) (combine q stags_ts); - ret (edb, searchRes). - - Definition G3 := G_gen Initialize_3. - - (* We need to show that the bad event is related to TSet correctness. Put the game in the correct form. *) - Definition TSetCorA := - [db, q, s_A] <-$3 A1; - k_S <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_2 db k_S k_I k_Z) (toW db); - [t_entries, sigmas] <-2 split wLoopRes; - t <- combine (toW db) t_entries; - ret (t, (fst (split q))). - - Theorem G2_G3_close : - | Pr[G2] - Pr[G3] | <= Pr[AdvCor _ TSetSetup TSetGetTag TSetRetrieve TSetCorA]. - - Abort. - - - (* Step 4: Give the plaintexts to GenTrans and remove the decryption.*) - Definition Initialize_indLoop_4 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_4 db k_S k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <- F k_S w; - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition ServerSearchLoop_4 xSet (e : nat * (Bvector (lambda + lambda) * Bvector lambda)) := - [xtoken, t_c_ind] <-2 e; - [t_c, ind] <-2 t_c_ind; - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some (e, ind)) else None. - - Definition ServerSearch_4 xSet (xtokens : list nat) (t : list (Bvector (lambda + lambda) * Bvector lambda)) := - map (ServerSearchLoop_4 xSet) (combine xtokens t). - - Definition GenTrans_4 (edb : EDB) k_X k_Z (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) : (list (Identifier lambda) * SearchTranscript) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <- map (fun (c : nat) => g^(e * (F_P k_Z (s ++ c)))) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - (inds, (stag, (combine xtokens res))). - - Definition Initialize_4 (db : DB)(q : list Query) := - k_S <-$ {0,1}^lambda; - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_4 db k_S k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G4 := G_gen Initialize_4. - - Theorem G3_G4_equiv : - Pr[G3] == Pr[G4]. - Abort. - - (* Step 5: prepare to the replace the first PRF with a random function. *) - Fixpoint oc_compMap(A B C D : Set)(eqdb : EqDec B)(c : A -> OracleComp C D B)(ls : list A) : OracleComp C D (list B) := - match ls with - | nil => $ (ret nil) - | a :: ls' => - b <--$ c a; - lsb' <--$ oc_compMap _ c ls'; - $ (ret (b :: lsb')) - end. - - Notation "'query' v" := (OC_Query _ v) (at level 79) : comp_scope. - - Definition Initialize_wLoop_5 db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$ query w; - indLoopRes <--$$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition Initialize_5 (db : DB)(q : list Query) := - - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - k_Z <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_5 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G5_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_5 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - Definition G5 := - k_S <-$ {0,1}^lambda; - [b, _] <-$2 G5_A _ _ (f_oracle F _ k_S) tt; - ret b. - - Theorem G4_G4_equiv : - Pr[G4] == Pr[G5]. - - Abort. - - (* Step 6: replace the PRF with a random function. *) - Definition G6 := - [b, _] <-$2 G5_A _ _ (@randomFunc Blist (Bvector lambda) (Rnd lambda) _) nil; - ret b. - - Theorem G5_G6_equiv : - | Pr[G5] - Pr[G6] | == PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ G5_A. - - Abort. - - (* Step 7: replace random function outputs with random values. *) - Definition Initialize_wLoop_7 db k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <-$ {0, 1}^lambda; - indLoopRes <-$ compMap _ (Initialize_indLoop_4 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_7 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_7 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G7 := G_gen Initialize_7. - - Theorem G6_G7_equiv : - Pr[G6] == Pr[G7]. - Abort. - - (* Step 8: Apply the semantic security definition to replace the ciphertexts with encryptions of 0^lambda *) - Definition zeroVector n := Vector.const false n. - - Definition Initialize_indLoop_8 k_I k_Z k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <-$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_8 db k_I k_Z w := - inds <- lookupInds db w; - sigma <-$ RndPerm (length inds); - k_E <-$ {0, 1}^lambda; - indLoopRes <-$ compMap _ (Initialize_indLoop_8 k_I k_Z k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - ret (indLoopRes, sigma). - - Definition Initialize_8 (db : DB)(q : list Query) := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - wLoopRes <-$ compMap _ (Initialize_wLoop_8 db k_I k_Z) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - ret (edb, searchRes). - - Definition G8 := G_gen Initialize_8. - - Require Import Encryption. - - Definition Initialize_indLoop_7_2 k_I k_Z w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$ query ind; - xind <- F_P k_I (Vector.to_list ind); - z <- F_P k_Z (w ++ c); - y <- xind * (inverse z); - $ ret (@Vector.append bool lambda lambda e (natToBvector y), ind). - - Definition Initialize_wLoop_7_2 db k_I k_Z w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_7_2 k_I k_Z w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition EncI_A1 := - [db, q, s_A] <-$3 A1; - k_I <-$ {0,1}^lambda; - k_Z <-$ {0,1}^lambda; - ret (toW db, (k_I, k_Z, db, q, s_A)). - - Definition EncI_A2 (s : Bvector lambda * Bvector lambda * DB * list Query * A_State) w := - let '(k_I, k_Z, db, q, s_A) := s in - Initialize_wLoop_7_2 db k_I k_Z w. - - Definition EncI_A3 (s : Bvector lambda * Bvector lambda * DB * list Query * A_State) wLoopRes := - let '(k_I, k_Z, db, q, s_A) := s in - k_X <-$ {0,1}^lambda; - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <-$2 TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <-$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <- map (GenTrans_4 edb k_X k_Z) (combine q stags_ts); - A2 s_A edb (snd (split searchRes)). - - Theorem G7_G8_equiv : - | Pr[G7] - Pr[G8] | <= IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc - EncI_A1 EncI_A2 EncI_A3 _ _ (oneVector lambda). - - Abort. - - (* Step 9: Next we will replace the "Z" PRF with a random function. Start by putting the game in the oracle form. *) - Definition Initialize_indLoop_9 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ query (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y), ind). - - Definition Initialize_wLoop_9 db k_I w := - inds <- lookupInds db w; - sigma <--$$ RndPerm (length inds); - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_9 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition GenTrans_9 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret g^(e * z)) (allNatsLt (length t)); - res <- ServerSearch_4 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine xtokens res))). - - Definition Initialize_9 (db : DB)(q : list Query) := - k_X <--$$ {0,1}^lambda; - k_I <--$$ {0,1}^lambda; - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_9 db k_I) (toW db); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - xSet <- XSetSetup_2 k_X k_I db sigmas; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_9 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G9_A := - [db, q, s_A] <--$3$ A1; - z0 <--$ Initialize_9 db q; - [edb, ls]<-2 z0; - $ A2 s_A edb (snd (split ls)). - - Definition G9 := - k <-$ {0, 1}^lambda; - [b, _] <-$2 G9_A _ _ (f_oracle F_P _ k) tt; - ret b. - - Theorem G8_G9_equiv : - Pr[G8] == Pr[G9]. - Abort. - - (* Step 10: Replace the PRF with a random function. *) - Definition G10 := - [b, _] <-$2 G9_A _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - ret b. - - Theorem G9_G10_quiv : - | Pr[G9] - Pr[G10] | <= PRF_Advantage (Rnd lambda) (RndNat (2^lambda)) F_P _ _ G9_A. - - Abort. - - - (* Step 11: Keep track of the blinding values and make sure nothing matches when we use the wrong value to unblind. *) - Definition Initialize_indLoop_11 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ query (w ++ c); - y <- xind * (inverse z); - $ ret (Vector.append e (natToBvector y), (ind, z)). - - Definition Initialize_wLoop_11 db k_I e := - [w, sigma] <-2 e; - inds <- lookupInds db w; - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_11 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition ServerSearchLoop_11 xSet (e : (nat * nat) * (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - [xtoken_z1, t_c_ind_z2] <-2 e; - [xtoken, z1] <-2 xtoken_z1; - [t_c, ind_z2] <-2 t_c_ind_z2; - [ind, z2] <-2 ind_z2; - if (eq_nat_dec z1 z2) then - ( - [e, y] <-2 splitVector lambda _ t_c; - if (in_dec (EqDec_dec _) (xtoken^(bvectorToNat y)) xSet) then (Some (e, ind)) else None - ) - else None. - - Definition ServerSearch_11 xSet (xtokens : list (nat * nat)) (t : list (Bvector (lambda + lambda) * (Bvector lambda * nat))) := - map (ServerSearchLoop_11 xSet) (combine xtokens t). - - Definition GenTrans_11 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda * nat)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret (g^(e * z), z)) (allNatsLt (length t)); - res <- ServerSearch_11 xSet xtokens t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine (fst (split xtokens)) res))). - - Definition Initialize_11 k_X k_I (db : DB)(q : list Query) sigmas xSet := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_11 db k_I) (combine (toW db) sigmas); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_11 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition sampleSigmas_11 db w := - inds <- lookupInds db w; - RndPerm (length inds). - - Definition G11 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, _] <-$2 (Initialize_11 k_X k_I db q sigmas xSet) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - [edb, ls] <-2 z; - A2 s_A edb (snd (split ls)). - - - Variable MaxMatchingIDs : nat. - Hypothesis MaxMatchingIDs_correct : - forall db q s_A, - In (db, q, s_A) (getSupport A1) -> - forall w, - (length (lookupInds db w) <= MaxMatchingIDs)%nat. - - Variable MaxKeywords : nat. - Hypothesis MaxKeywords_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length (toW db) <= MaxKeywords)%nat. - Variable MaxQueries : nat. - Hypothesis MaxQueries_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length q <= MaxQueries)%nat. - - Variable MaxIDs : nat. - Hypothesis MaxIDs_correct : - forall db q s, In (db, q, s) (getSupport A1) -> - (length db <= MaxIDs)%nat. - - Theorem G10_G11_equiv : - | Pr[G10] - Pr[G11] | <= (((MaxKeywords + MaxQueries) * MaxMatchingIDs)^2) * - MaxQueries * MaxMatchingIDs * MaxIDs * - MaxKeywords * MaxMatchingIDs / 2 ^ lambda. - Abort. - - (* Setp 12 : Add some bookkeeping so we can compute the transcript differently and replace the random function outputs with random values. *) - - - (* Step 12: Remove some of the query blinding and simplify. *) - Definition Initialize_indLoop_12 k_I k_E w (e : Identifier lambda * nat) := - [ind, c] <-2 e; - e <--$$ Enc k_E (zeroVector lambda); - xind <- F_P k_I (Vector.to_list ind); - z <--$ OC_Query nat (w ++ c); - $ ret (Vector.append e (natToBvector xind), (ind, z)). - - Definition Initialize_wLoop_12 db k_I e := - [w, sigma] <-2 e; - inds <- lookupInds db w; - k_E <--$$ {0, 1}^lambda; - indLoopRes <--$ oc_compMap _ (Initialize_indLoop_12 k_I k_E w) (combine (permute inds sigma) (allNatsLt (length inds))); - $ ret (indLoopRes, sigma). - - Definition GenTrans_12 (edb : EDB) k_X (e : Query * (Tag * list (Bvector (lambda + lambda) * (Bvector lambda * nat)))) := - [q, stag_t] <-2 e; - [stag, t] <-2 stag_t; - [s, x] <-2 q; - [tSet, xSet] <-2 edb; - e <- F_P k_X x; - xtokens <--$ oc_compMap _ (fun (c : nat) => z <--$ query s ++ c; $ ret (g^(e * z), (g ^ e, z))) (allNatsLt (length t)); - res <- ServerSearch_11 xSet (snd (split xtokens)) t; - es <- getSomes res; - res <- map (fun z => match z with - | Some y => Some (fst y) - | None => None - end) res; - inds <- map (fun x => snd x) es; - $ ret (inds, (stag, (combine (fst (split xtokens)) res))). - - Definition Initialize_12 k_X k_I (db : DB)(q : list Query) sigmas xSet := - wLoopRes <--$ oc_compMap _ (Initialize_wLoop_11 db k_I) (combine (toW db) sigmas); - [t_entries_pts, sigmas] <-2 split wLoopRes; - t_entries <- map (fun v => (fst (split v))) t_entries_pts; - t <- combine (toW db) t_entries; - [tSet, k_T] <--$2$ TSetSetup t; - t <- combine (toW db) t_entries_pts; - edb <- (tSet, xSet); - stags_ts <--$$ compMap _ (fun x => tag <-$ TSetGetTag k_T x; ret (tag, (arrayLookupList _ t x))) (fst (split q)); - searchRes <--$ oc_compMap _ (GenTrans_12 edb k_X) (combine q stags_ts); - $ ret (edb, searchRes). - - Definition G12 := - k_X <-$ {0,1}^lambda; - k_I <-$ {0,1}^lambda; - [db, q, s_A] <-$3 A1; - sigmas <-$ compMap _ (sampleSigmas_11 db) (toW db); - xSet <- XSetSetup_2 k_X k_I db sigmas; - [z, _] <-$2 (Initialize_12 k_X k_I db q sigmas xSet) _ _ (@randomFunc Blist nat (RndNat (2^lambda)) _ ) nil; - [edb, ls] <-2 z; - A2 s_A edb (snd (split ls)). - -End ESPADA_SSE_OXT_Games. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_SSE_SKS.v b/src/ESPADA/ESPADA_SSE_SKS.v deleted file mode 100644 index f5be956..0000000 --- a/src/ESPADA/ESPADA_SSE_SKS.v +++ /dev/null @@ -1,260 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import SSE. -Require Import TSet. -Require Import RndPerm. -Require Import CompFold. -Require Import Array. -Require Import PRF. -Require Import Encryption. -Require Import OracleCompFold. - -Local Open Scope list_scope. - -Section ESPADA_SSE_SKS. - - Variable lambda : posnat. - Definition DB := DB lambda. - - (* PRF *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - - (* TSet *) - Variable TSet : Set. - Hypothesis TSet_EqDec : EqDec TSet. - Variable TSetKey : Set. - Hypothesis TSetKey_EqDec : EqDec TSetKey. - Variable Tag : Set. - Hypothesis Tag_EqDec : EqDec Tag. - Variable TSetSetup : T (pos (lambda)) -> Comp (TSet * TSetKey). - Variable TSetGetTag : TSetKey -> Keyword -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector (lambda))). - Variable Leakage_T : Set. - Hypothesis Leakage_T_EqDec : EqDec Leakage_T. - Variable L_T : T (pos (lambda)) -> list Keyword -> Leakage_T. - Variable Sim_T : Leakage_T -> list (list (Bvector lambda)) -> Comp (TSet * list Tag). - - Definition lookupInds(db : DB)(w : Keyword) := - fold_left (fun ls p => if (in_dec (EqDec_dec _) w (snd p)) then (fst p :: ls) else ls) db nil. - - Definition toW (db : DB) := - removeDups _ (flatten (snd (split db))). - - Definition SKS_EDBSetup_indLoop k_e ind := - Enc k_e ind. - - Definition SKS_EDBSetup_wLoop db k_S w := - k_e <- F k_S w; - inds <- lookupInds db w; - t <-$ compMap _ (SKS_EDBSetup_indLoop k_e) inds; - ret (w, t). - - Definition SKS_EDBSetup(db : DB) := - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - ret ((k_S, k_T), tSet). - - Definition SKS_Search tSet k w := - [k_S, k_T] <-2 k; - (* client *) - tag <-$ TSetGetTag k_T w; - (* server *) - t <- TSetRetrieve tSet tag; - (* client *) - k_e <- F k_S w; - inds <- map (Dec k_e) t; - (* Search returns the results and the transcript. *) - ret (inds, (tag, t)). - - Definition EDB := TSet. - Definition SearchTranscript := (Tag * list (Bvector lambda))%type. - Definition Query := Keyword. - Definition Identifier := Bvector lambda. - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - Definition zeroVector n := Vector.const false n. - - Definition SKS_resultsStruct db w := - k_e <-$ (Rnd lambda); - inds <- lookupInds db w; - compMap _ (fun _ => Enc k_e (zeroVector lambda)) inds. - - Definition L (db : DB) (qs : list Query) := - t_struct <-$ compMap _ (SKS_resultsStruct db) (toW db); - t <- combine (toW db) t_struct; - leak_T <- L_T t qs; - ret (leak_T, map (arrayLookupList _ t) qs). - - Definition SKS_Sim (leak : _ * list (list (Bvector lambda))) := - [leak_T, resStruct] <-2 leak; - [tSet, tags] <-$2 Sim_T leak_T resStruct; - ret (tSet, (combine tags resStruct)). - - - (* Step 1: inline and simplify *) - Definition SKS_Search_1 tSet k_T w := - tag <-$ TSetGetTag k_T w; - t <- TSetRetrieve tSet tag; - ret (tag, t). - - Definition G1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - ls <-$ compMap _ (SKS_Search_1 tSet k_T) q; - A2 s_A tSet ls. - - - (* Step 2: factor out some TSet computations and prepare to apply TSet correctness*) - - Definition G2 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - ls <- combine tags rs; - A2 s_A tSet ls. - - - (* Step 3: apply TSet correctness. *) - Definition G3 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition TSetCor_A := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - ret (t, q). - - - (* Step 4: apply TSet security. *) - Definition G4 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition TSetSec_A1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - rs <- map (arrayLookupList _ t) q; - ret (t, q, (s_A, rs)). - - Definition TSetSec_A2 s p := - [s_A, rs] <-2 s; - [tSet, tags] <-2 p; - ls <- combine tags rs; - A2 s_A tSet ls. - - (* Step 5-6: apply PRF definition. *) - Definition SKS_EDBSetup_wLoop_5 db w := - k_e <--$ query w; - inds <- lookupInds db w; - t <--$$ compMap _ (SKS_EDBSetup_indLoop k_e) inds; - $ ret (w, t). - - Definition PRF_A := - [db, q, s_A] <--$3$ A1; - t <--$ oc_compMap _ (SKS_EDBSetup_wLoop_5 db) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <--$2$ Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - $ A2 s_A tSet ls. - - Definition G5 := - k_S <-$ {0, 1}^lambda; - [b, _] <-$2 PRF_A _ _ (f_oracle F _ k_S) tt; - ret b. - - - Definition G6 := - [b, _] <-$2 PRF_A _ _ (RndR_func (Rnd lambda) _) nil; - ret b. - - - (* Step 7: replace RF outputs with random values. *) - Definition SKS_EDBSetup_wLoop_7 db w := - k_e <-$ (Rnd lambda); - inds <- lookupInds db w; - t <-$ compMap _ (SKS_EDBSetup_indLoop k_e) inds; - ret (w, t). - - Definition G7 := - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop_7 db) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - (* Step 8: apply encryption security definition. *) - Definition G8 := - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_resultsStruct db) (toW db); - t <- combine (toW db) t; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition Enc_A1 := - [db, q, s_A] <-$3 A1; - ret ((map (pair db) (toW db)), (db, q, s_A)). - - Definition Enc_A2 e := - [db, w] <-2 e; - oc_compMap _ (@OC_Query _ (Bvector lambda)) (lookupInds db w). - - Definition Enc_A3 s_A t := - [db, q, s_A] <-3 s_A; - t <- combine (toW db) t; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition G7_G8_Distance := - IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc Enc_A1 Enc_A2 Enc_A3 _ _(zeroVector lambda). - - (* G8 is equal to the ideal game*) - - - -End ESPADA_SSE_SKS. - - diff --git a/src/ESPADA/ESPADA_SSE_SKS_Secure.v b/src/ESPADA/ESPADA_SSE_SKS_Secure.v deleted file mode 100644 index eb6a793..0000000 --- a/src/ESPADA/ESPADA_SSE_SKS_Secure.v +++ /dev/null @@ -1,1205 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import SSE. -Require Import TSet. -Require Import RndPerm. -Require Import CompFold. -Require Import Array. -Require Import PRF. -Require Import Encryption. -Require Import OracleCompFold. -Require Import ESPADA_SSE_SKS. - -Local Open Scope list_scope. - - -Section ESPADA_SSE_SKS_Secure. - - Variable lambda : posnat. - - (* PRF *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - - (* TSet *) - Variable TSet : Set. - Hypothesis TSet_EqDec : EqDec TSet. - Variable TSetKey : Set. - Hypothesis TSetKey_EqDec : EqDec TSetKey. - Variable Tag : Set. - Hypothesis Tag_EqDec : EqDec Tag. - Variable TSetSetup : T (pos (lambda)) -> Comp (TSet * TSetKey). - Variable TSetGetTag : TSetKey -> Keyword -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector (lambda))). - Variable Leakage_T : Set. - Hypothesis Leakage_T_EqDec : EqDec Leakage_T. - Variable L_T : T (pos (lambda)) -> list Keyword -> Leakage_T. - Variable Sim_T : Leakage_T -> list (list (Bvector lambda)) -> Comp (TSet * list Tag). - - Definition EDB := EDB TSet. - Definition SearchTranscript := SearchTranscript lambda Tag. - Definition SKS_EDBSetup := SKS_EDBSetup F Enc _ _ TSetSetup. - Definition SKS_Search := @SKS_Search lambda F Dec TSet TSetKey Tag _ - TSetGetTag TSetRetrieve. - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB lambda * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - - Notation "'gInst' g" := (g lambda F Enc TSet TSetKey Tag _ TSetSetup TSetGetTag TSetRetrieve _ A1 A2) (at level 90). - - (* Step 1: inline and simplify *) - Theorem Real_G1_equiv : - Pr[SSE_Sec_NA_Real SKS_EDBSetup _ SKS_Search A1 A2] == Pr[gInst G1]. - - unfold G1, SSE_Sec_NA_Real, SKS_EDBSetup. - do 4 (inline_first; comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply (@compMap_spec _ _ _ _ _ _ eq (fun a b => snd a = b)). - eapply list_pred_eq. - intuition; subst. - unfold SKS_Search, SKS_Search_1. - comp_skip. - eapply comp_base_exists; intuition. - eapply comp_base_exists; intuition. - eapply comp_spec_ret; intuition. - erewrite snd_split_eq_list_pred. - eapply comp_spec_eq_refl. - intuition. - - Qed. - - (* Step 2: factor out some TSet computations and prepare to apply TSet correctness*) - - Definition SKS_EDBSetup_wLoop := SKS_EDBSetup_wLoop F Enc. - - Definition G1_1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - ls <-$ ( - tags <-$ compMap _ (TSetGetTag k_T) q; - ret (map (fun z => (z, TSetRetrieve tSet z)) tags) - ); - A2 s_A tSet ls. - - - Theorem G1_G1_1_equiv : - Pr[gInst G1] == Pr[G1_1]. - - unfold G1, G1_1, SKS_EDBSetup_wLoop. - do 4 (comp_skip; comp_simp). - comp_skip. - symmetry. - eapply comp_spec_eq_impl_eq. - - eapply compMap_map_fission_eq; intuition. - eapply comp_spec_consequence. - eapply comp_spec_eq_refl. - intuition; subst. - trivial. - - Qed. - - Theorem G1_1_G2_equiv : - Pr[G1_1] == Pr[gInst G2]. - - unfold G1_1, G2, SKS_EDBSetup_wLoop. - do 4 (comp_skip; comp_simp). - inline_first. - comp_skip. - rewrite combine_map_eq. - eapply eqRat_refl_eq. - f_equal. - f_equal. - eapply (@map_ext_pred _ _ _ (fun a b => a = fst b /\ a = snd b)). - eapply list_pred_symm. - eapply list_pred_impl. - eapply (@list_pred_combine_l _ _ _ (fun a b => a = b) (fun a b => a = b)). - eapply list_pred_eq. - eapply list_pred_eq. - intuition. - - intuition; subst. - rewrite H8. - trivial. - - Qed. - - Theorem G1_G2_equiv : - Pr[gInst G1] == Pr[gInst G2]. - - rewrite G1_G1_1_equiv. - eapply G1_1_G2_equiv. - - Qed. - - (* Step 3: apply TSet correctness. *) - Definition G2_1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - v <-$ A2 s_A tSet ls; - ret (v, negb (eqb rs rs')). - - Definition G2_2 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ls <- combine tags rs'; - v <-$ A2 s_A tSet ls; - ret (v, negb (eqb rs rs')). - - Theorem G2_G2_1_equiv : - Pr[gInst G2] == Pr[x <-$ G2_1; ret fst x]. - - unfold G2, G2_1, SKS_EDBSetup_wLoop. - repeat (inline_first; comp_skip; comp_simp). - inline_first. - rewrite <- evalDist_right_ident. - comp_skip. - comp_simp. - simpl; intuition. - - Qed. - - Theorem G2_1_G2_2_eq_until_bad : - forall x, - evalDist G2_1 (x, false) == evalDist G2_2 (x, false). - - intuition. - unfold G2_1, G2_2. - do 5 (comp_skip; comp_simp). - case_eq ((@eqb (list (list (Bvector (posnatToNat lambda)))) - (@list_EqDec (list (Bvector (posnatToNat lambda))) - (@list_EqDec (Bvector (posnatToNat lambda)) - (Bvector_EqDec (posnatToNat lambda)))) - (@map Tag (list (Bvector (posnatToNat lambda))) - (TSetRetrieve t) x2) - (@map Keyword (list (Bvector (posnatToNat lambda))) - (@arrayLookupList Keyword - (Bvector (posnatToNat lambda)) - (@list_EqDec bool bool_EqDec) x1) l))); intuition. - rewrite eqb_leibniz in H6. - rewrite H6. - eapply eqRat_refl. - - comp_irr_l. - comp_irr_r. - - repeat rewrite evalDist_ret_0; intuition. - pairInv. - pairInv. - - Qed. - - Theorem G2_1_G2_2_bad_same : - Pr[x <-$ G2_1; ret snd x] == Pr[x <-$ G2_2; ret snd x]. - - unfold G2_1, G2_2. - do 5 (inline_first; comp_skip; comp_simp). - inline_first. - comp_irr_l. - comp_irr_r. - comp_simp. - simpl. - intuition. - - Qed. - - Definition G2_1_bad := - [t, q] <-$2 (TSetCor_A F Enc A1); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ret (negb (eqb rs rs')). - - Theorem G2_1_bad_equiv : - Pr[x <-$ G2_1; ret snd x] == Pr[G2_1_bad]. - - unfold G2_1, G2_1_bad, TSetCor_A. - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - inline_first. - comp_irr_l. - simpl; intuition. - - Qed. - - Theorem G2_1_bad_small : - Pr[G2_1_bad] <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - eapply leRat_refl. - - Qed. - - Theorem G2_1_G2_2_close : - | Pr[x <-$ G2_1; ret fst x] - Pr[x <-$ G2_2; ret fst x] | <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - intuition. - eapply leRat_trans. - eapply fundamental_lemma_h. - eapply G2_1_G2_2_bad_same. - eapply G2_1_G2_2_eq_until_bad. - rewrite G2_1_bad_equiv. - eapply G2_1_bad_small. - - Qed. - - Notation "'gInst1' g" := (g lambda F Enc TSet TSetKey Tag _ TSetSetup TSetGetTag _ A1 A2) (at level 90). - - Theorem G2_2_G3_equiv : - Pr[x <-$ G2_2; ret fst x] == Pr[gInst1 G3]. - - unfold G2_2, G3. - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - inline_first. - symmetry. - rewrite <- evalDist_right_ident. - comp_skip. - comp_simp. - simpl; intuition. - - Qed. - - Theorem G2_G3_equiv : - | Pr[gInst G2] - Pr[gInst1 G3] | <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - rewrite G2_G2_1_equiv. - rewrite <- G2_2_G3_equiv. - eapply G2_1_G2_2_close. - - Qed. - - - (* Step 4: apply TSet security. *) - Definition TSetSec_A1 := TSetSec_A1 F Enc _ A1. - Definition TSetSec_A2 := TSetSec_A2 A2. - Definition G3_1 := - [t, q, s] <-$3 TSetSec_A1; - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - TSetSec_A2 s (tSet, tags). - - Definition G3_2 := - [t, q, s] <-$3 TSetSec_A1; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - TSetSec_A2 s (tSet, tags). - - Theorem G3_G3_1_equiv : - Pr[gInst1 G3] == Pr[G3_1]. - - unfold G3, G3_1, TSetSec_A1, TSetSec_A2. - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - - Qed. - - Theorem G3_1_G3_2_close: - | Pr[G3_1] - Pr[G3_2] | <= TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T. - - eapply leRat_refl. - - Qed. - - Notation "'gInst2' g" := (g lambda F Enc TSet Tag _ L_T Sim_T _ A1 A2) (at level 90). - - Theorem G3_2_G4_equiv : - Pr[G3_2] == Pr[gInst2 G4]. - - unfold G3_2, G4, TSetSec_A1, TSetSec_A2. - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - Qed. - - Theorem G3_G4_equiv : - | Pr[gInst1 G3] - Pr[gInst2 G4] | <= TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T. - - rewrite G3_G3_1_equiv. - rewrite <- G3_2_G4_equiv. - eapply G3_1_G3_2_close. - - Qed. - - (* Step 5-6: apply PRF definition. *) - Definition G4_1 := - k_S <-$ {0, 1}^lambda; - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition SKS_EDBSetup_wLoop_5 := SKS_EDBSetup_wLoop_5 Enc. - - Theorem SKS_EDBSetup_5_spec : - forall d b z x, - comp_spec - (fun (x0 : Keyword * list (Bvector lambda)) - (y : Keyword * list (Bvector lambda) * unit) => - x0 = (fst y)) (SKS_EDBSetup_wLoop d x b) - ((SKS_EDBSetup_wLoop_5 d b) unit unit_EqDec - (f_oracle F (Bvector_EqDec lambda) x) z). - - intuition. - unfold SKS_EDBSetup_wLoop, SKS_EDBSetup_wLoop_5. - simpl. - unfold f_oracle. - comp_simp. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G4_G4_1_equiv : - Pr[gInst2 G4] == Pr[G4_1]. - - unfold G4, G4_1. - comp_swap_r. - repeat (comp_skip; comp_simp). - reflexivity. - Qed. - - Theorem G4_1_G5_equiv : - Pr[G4_1] == Pr[gInst2 G5]. - - Local Opaque evalDist. - - unfold G4_1, G5, PRF_A. - repeat (inline_first; comp_skip; comp_simp; simpl). - inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_oc_spec. - eapply list_pred_eq. - intuition; subst. - eapply SKS_EDBSetup_5_spec. - - comp_simp. - simpl in H4. - eapply list_pred_eq_impl_eq in H4. - subst. - inline_first. - comp_skip. - comp_simp. - simpl. - inline_first. - eapply comp_spec_eq_trans. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - - Qed. - - Theorem G4_G5_equiv : - Pr[gInst2 G4] == Pr[gInst2 G5]. - - rewrite G4_G4_1_equiv. - eapply G4_1_G5_equiv. - - Qed. - - Notation "'gInst3' g" := (g lambda Enc TSet Tag _ L_T Sim_T _ A1 A2) (at level 90). - Definition PRF_A := PRF_A Enc L_T Sim_T A1 A2. - - Theorem G5_G6_equiv : - | Pr[gInst2 G5] - Pr[gInst3 G6] | <= PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ PRF_A. - - eapply leRat_refl. - - Qed. - - Definition SKS_EDBSetup_wLoop_6_f db w k_e := - inds <- lookupInds db w; - t <-$ compMap _ (SKS_EDBSetup_indLoop _ Enc k_e) inds; - ret (w, t). - - Definition SKS_EDBSetup_wLoop_6 db w := - k_e <--$ query w; - $ (SKS_EDBSetup_wLoop_6_f db w k_e). - - Definition PRF_A_6 := - [db, q, s_A] <--$3$ A1; - t <--$ oc_compMap _ (SKS_EDBSetup_wLoop_6 db) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <--$2$ Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - $ A2 s_A tSet ls. - - Definition G6_1 := - [b, _] <-$2 PRF_A_6 _ _ (RndR_func (Rnd lambda) _) nil; - ret b. - - - Theorem G6_G6_1_equiv : - Pr[gInst3 G6] == Pr[G6_1]. - - unfold G6, G6_1. - - unfold PRF_A, PRF_A_6. - do 2 (simpl; inline_first; comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply oc_compMap_eq . - intuition. - unfold SKS_EDBSetup_wLoop_5, SKS_EDBSetup_wLoop_6. - simpl. - comp_skip. - unfold SKS_EDBSetup_wLoop_6_f. - comp_simp. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - Qed. - - Definition G6_2 := - [b, _] <-$2 PRF_A_6 _ _ (fun s a => x <-$ Rnd lambda; ret (x, s)) tt; - ret b. - - Theorem G6_1_G6_2_equiv : - Pr[G6_1] == Pr[G6_2]. - - unfold G6_1, G6_2, PRF_A_6. - simpl; inline_first. - comp_skip. - comp_simp. - simpl; inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_randomFunc_NoDup. - unfold toW. - eapply removeDups_NoDup. - intuition. - - inline_first. - comp_simp. - inline_first. - simpl in *. - subst. - comp_skip. - simpl; inline_first. - comp_simp. - simpl; inline_first. - comp_skip. - comp_simp. - - eapply comp_spec_eq_refl. - - Qed. - - - (* Step 7: replace RF outputs with random values. *) - - Theorem EDBSetup_7_spec : - forall d b z, - comp_spec - (fun (x : Keyword * list (Bvector lambda)) - (y : Keyword * list (Bvector lambda) * unit) => - x = (fst y)) (SKS_EDBSetup_wLoop_7 Enc d b) - ((SKS_EDBSetup_wLoop_6 d b) unit unit_EqDec - (fun (s : unit) (_ : Keyword) => x <-$ { 0 , 1 }^lambda; ret (x, s)) - z). - - intuition. - unfold SKS_EDBSetup_wLoop_6, SKS_EDBSetup_wLoop_7. - simpl. - inline_first. - comp_skip. - comp_simp. - unfold SKS_EDBSetup_wLoop_6_f. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G6_2_G7_equiv : - Pr[G6_2] == Pr[gInst3 G7]. - - unfold G6_2, G7. - unfold PRF_A_6. - simpl; inline_first. - comp_skip. - comp_simp. - simpl; inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply comp_spec_symm. - eapply compMap_oc_spec. - eapply list_pred_eq. - intuition; subst. - eapply EDBSetup_7_spec. - inline_first. - simpl in *. - eapply list_pred_eq_impl_eq in H2. - subst. - comp_skip. - simpl; inline_first. - eapply comp_spec_eq_trans. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - - Qed. - - Theorem G6_G7_equiv : - Pr[gInst3 G6] == Pr[gInst3 G7]. - - rewrite G6_G6_1_equiv. - rewrite G6_1_G6_2_equiv. - eapply G6_2_G7_equiv. - - Qed. - - (* Step 8: apply encryption security definition. *) - Definition SKS_EDBSetup_wLoop_7_1 db w := - k_e <-$ (Rnd lambda); - inds <- lookupInds db w; - compMap _ (SKS_EDBSetup_indLoop _ Enc k_e) inds. - - Definition G7_1 := - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_1 db) (toW db); - t <- combine (toW db) t; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Theorem EDBSetup_7_1_spec : - forall d b, - comp_spec (fun a b => snd a = b) (SKS_EDBSetup_wLoop_7 Enc d b) (SKS_EDBSetup_wLoop_7_1 d b). - - intuition. - unfold SKS_EDBSetup_wLoop_7, SKS_EDBSetup_wLoop_7_1. - comp_skip. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G7_G7_1_equiv : - Pr[gInst3 G7] == Pr[G7_1]. - - unfold G7, G7_1. - comp_skip; comp_simp. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition; subst. - eapply EDBSetup_7_1_spec. - eapply (@compMap_support _ _ (fun a b => fst b = a)) in H0. - assert (a0 = (combine (toW d) b)). - - apply snd_split_eq_list_pred in H2. - subst. - eapply list_pred_symm in H0. - eapply fst_split_eq_list_pred in H0. - rewrite <- H0. - symmetry. - specialize (split_combine a0); intuition. - remember (split a0) as z. - destruct z. - trivial. - - subst. - comp_skip. - - intuition. - unfold SKS_EDBSetup_wLoop_7 in *. - repeat simp_in_support. - trivial. - - Qed. - - - Definition SKS_EDBSetup_wLoop_7_2 w := - k_e <-$ (Rnd lambda); - [b, _] <-$2 (Enc_A2 w) _ _ (EncryptOracle Enc _ k_e) tt; - ret b. - - Definition G7_2 := - [lsa, s_A] <-$2 (Enc_A1 _ A1); - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_2) lsa; - Enc_A3 L_T Sim_T A2 s_A t. - - Theorem EncOracle_spec : - forall b0 z x1, - comp_spec - (fun (x2 : Bvector lambda) (y : Bvector lambda * unit) => - x2 = (fst y)) (Enc x1 b0) - ((query b0) unit unit_EqDec - (EncryptOracle Enc (Bvector_EqDec lambda) x1) z). - - intuition. - unfold EncryptOracle. - simpl. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G7_1_G7_2_equiv : - Pr[G7_1] == Pr[G7_2]. - - unfold G7_1, G7_2. - unfold Enc_A1. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply (@compMap_eq _ _ (fun a b => fst b = d /\ a = snd b)). - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - intuition; subst. - unfold SKS_EDBSetup_wLoop_7_1, SKS_EDBSetup_wLoop_7_2. - comp_skip. - unfold Enc_A2. - rewrite <- evalDist_right_ident. - eapply comp_spec_eq_impl_eq. - destruct b; simpl. - comp_skip. - eapply compMap_oc_spec. - eapply list_pred_eq. - intuition; subst. - unfold SKS_EDBSetup_indLoop. - eapply EncOracle_spec. - comp_simp. - eapply comp_spec_ret; intuition. - simpl in *. - eapply list_pred_eq_impl_eq; intuition. - - unfold Enc_A3. - comp_simp. - eapply eqRat_refl. - - Qed. - - Definition SKS_EDBSetup_wLoop_7_3 w := - k_e <-$ (Rnd lambda); - [b, _] <-$2 (Enc_A2 w) _ _ (EncryptNothingOracle Enc _ (zeroVector lambda) k_e) tt; - ret b. - - Definition G7_3 := - [lsa, s_A] <-$2 Enc_A1 _ A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_3) lsa; - Enc_A3 L_T Sim_T A2 s_A t. - - Theorem G7_2_G7_3_close : - | Pr[G7_2] - Pr[G7_3] | <= - IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc (Enc_A1 _ A1) (@Enc_A2 lambda) (Enc_A3 L_T Sim_T A2) _ _(zeroVector lambda). - - eapply leRat_refl. - Qed. - - - Theorem EncNothing_spec : - forall b0 z x1, - comp_spec - (fun (x2 : Bvector lambda) (y : Bvector lambda * unit) => - x2 = (fst y)) (Enc x1 (zeroVector lambda)) - ((query b0) unit unit_EqDec - (EncryptNothingOracle Enc (Bvector_EqDec lambda) - (zeroVector lambda) x1) z). - - intuition. - unfold EncryptNothingOracle. - simpl. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip. - eapply comp_spec_ret; intuition. - Qed. - - - Notation "'gInst4' g" := (g _ Enc _ _ _ L_T Sim_T _ A1 A2) (at level 90). - - Theorem G7_3_G8_equiv : - Pr[G7_3] == Pr[gInst4 G8]. - - unfold G7_3, G8. - unfold Enc_A1. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply (@compMap_eq _ _ (fun a b => fst a = d /\ b = snd a)). - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - intuition; subst. - unfold SKS_EDBSetup_wLoop_7_3, SKS_resultsStruct. - comp_skip. - unfold Enc_A2. - symmetry. - rewrite <- evalDist_right_ident. - eapply comp_spec_eq_impl_eq. - simpl. - comp_skip. - eapply compMap_oc_spec. - eapply list_pred_eq. - intuition; subst. - - eapply EncNothing_spec. - comp_simp. - simpl in *. - eapply comp_spec_ret. - eapply list_pred_eq_impl_eq; intuition. - - unfold Enc_A3. - comp_simp. - eapply eqRat_refl. - - Qed. - - Theorem G7_G8_equiv : - | Pr[gInst3 G7] - Pr[gInst4 G8] | <= - IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc (Enc_A1 _ A1) (@Enc_A2 lambda) (Enc_A3 L_T Sim_T A2) _ _(zeroVector lambda). - - rewrite G7_G7_1_equiv. - rewrite G7_1_G7_2_equiv. - rewrite <- G7_3_G8_equiv. - eapply G7_2_G7_3_close. - - Qed. - - - Theorem G8_Ideal_equiv : - Pr[gInst4 G8] == Pr[SSE_Sec_NA_Ideal A1 A2 (L Enc _ L_T) (SKS_Sim _ _ _ Sim_T)]. - - Local Opaque evalDist. - - unfold G8, SSE_Sec_NA_Ideal, L, SKS_Sim. - comp_skip. - comp_simp. - inline_first. - comp_skip. - comp_simp. - simpl. - inline_first. - comp_skip. - - reflexivity. - - comp_simp. - eapply eqRat_refl. - - Qed. - - Variable IND_CPA_Adv : Rat. - - Hypothesis IND_CPA_Adv_correct : - (forall i, - IND_CPA_SecretKey_O_Advantage ({0,1}^lambda) Enc - (A_c - (B1 (nil, nil) _ _ (Enc_A1 _ A1) i) - (@Enc_A2 lambda) - (B2 - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) _ _ (EncryptOracle Enc _ key) tt; - [b, _]<-2 z; ret b) - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) unit unit_EqDec - (EncryptNothingOracle Enc _ (zeroVector lambda) key) tt; - [b, _]<-2 z; ret b) _ (Enc_A3 L_T Sim_T A2)) - ) - _ - (zeroVector lambda) - <= IND_CPA_Adv)%rat. - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - forall db q s, - In (db, q, s) (getSupport A1) -> - (length (toW db) <= maxKeywords)%nat. - - Hypothesis Enc_wf : - forall k p, - well_formed_comp (Enc k p). - - Theorem ESPADA_SKS_Secure_concrete : - SSE_NA_Advantage SKS_EDBSetup _ SKS_Search A1 A2 (L Enc _ L_T) (SKS_Sim _ _ _ Sim_T) <= - AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1) + - TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T + - PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ PRF_A + - (maxKeywords / 1) * IND_CPA_Adv. - - unfold SSE_NA_Advantage. - rewrite Real_G1_equiv. - rewrite G1_G2_equiv. - eapply leRat_trans. - eapply ratDistance_le_trans. - eapply G2_G3_equiv. - eapply ratDistance_le_trans. - eapply G3_G4_equiv. - rewrite G4_G5_equiv. - eapply ratDistance_le_trans. - eapply G5_G6_equiv. - rewrite G6_G7_equiv. - rewrite <- G8_Ideal_equiv. - eapply G7_G8_equiv. - repeat rewrite <- ratAdd_assoc. - - eapply ratAdd_leRat_compat; intuition. - eapply IND_CPA_OI_impl_O. - intuition. - wftac. - intuition. - unfold Enc_A2. - - eapply oc_compMap_wf. - intuition. - econstructor. - intuition. - unfold Enc_A1 in *. - repeat simp_in_support. - destruct x. - destruct p. - repeat simp_in_support. - rewrite map_length. - eauto. - - eauto. - - intuition. - - Qed. - - Print Assumptions ESPADA_SKS_Secure_concrete. - -End ESPADA_SSE_SKS_Secure. - - -(* Just to make sure everything fits together, let's combine the result above with the results about the T-Set instantiation. This is just a sanity check to make sure everything unifies and we can derive the bounds. There is a lot of room for improvement in the proof below. *) -Require Import ESPADA_TSet. -Require Import ESPADA_TSet_Secure. -Require Import ESPADA_TSet_Correct. -Require Import ESPADA_TSet_Once_Games. -Require Import ESPADA_TSet_Once. -Require Import ESPADA_TSet_Correct_Once_Games. - -Section ESPADA_SSE_SKS_Secure_inst. - - Variable lambda : posnat. - - (* PRF *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Hypothesis Enc_wf : - forall k p, - well_formed_comp (Enc k p). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - - Variable bigB bigS : posnat. - Variable F_T : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (Datatypes.S lambda). - Variable F_bar_T : Bvector lambda -> Blist -> Bvector lambda. - Variable maxMatches maxKeywords : posnat. - - Definition TSetSetup := @ESPADA_TSetSetup lambda bigB bigS F_T F_bar_T. - Definition TSetGetTag := ESPADA_TSetGetTag lambda F_bar_T. - Definition TSetRetrieve := ESPADA_TSetRetrieve lambda maxMatches F_T. - - Definition TSetKey := Bvector lambda. - - Definition TSet := TSet lambda. - Definition EDB_T := EDB TSet. - Definition Tag := Bvector lambda. - Definition SearchTranscript_T := SearchTranscript lambda Tag. - Definition SKS_EDBSetup_T := SKS_EDBSetup F Enc _ _ TSetSetup. - Definition SKS_Search_T := @SKS_Search lambda F Dec TSet TSetKey Tag _ - TSetGetTag TSetRetrieve. - - Definition Sim_T := @ESPADA_TSet_Sim lambda bigB bigS F_T. - Definition L_T := @L_T lambda. - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB lambda * list Query * A_State). - Variable A2 : A_State -> EDB_T -> list SearchTranscript_T -> Comp bool. - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - Variable IND_CPA_Adv : Rat. - Hypothesis IND_CPA_Adv_correct : - (forall i, - IND_CPA_SecretKey_O_Advantage ({0,1}^lambda) Enc - (A_c - (B1 (nil, nil) _ _ (Enc_A1 _ A1) i) - (@Enc_A2 lambda) - (B2 - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) _ _ (EncryptOracle Enc _ key) tt; - [b, _]<-2 z; ret b) - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) unit unit_EqDec - (EncryptNothingOracle Enc _ (zeroVector lambda) key) tt; - [b, _]<-2 z; ret b) _ (Enc_A3 L_T Sim_T A2)) - ) - _ - (zeroVector lambda) - <= IND_CPA_Adv)%rat. - - Variable PRF_NA_Adv failProb : Rat. - Variable numTrials : posnat. - - Hypothesis failProb_pos : - ~(1 <= failProb). - - Hypothesis failProb_correct_Cor : - forall (t : T lambda) (l : list Blist), - In (t, l) (getSupport (TSetCor_A F Enc A1)) -> - Pr - [x <-$ ESPADA_TSet_Correct.ESPADA_TSetSetup_trial bigB bigS F_T F_bar_T t; - ret (if fst x then false else true) ] <= failProb. - - - Definition PRF_A_T := @PRF_A lambda Enc TSet Tag _ L_T Sim_T _ A1 A2. - - Hypothesis PRF_NA_Adv_correct_1 : - PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar_T - (list_EqDec bool_EqDec) (Bvector_EqDec lambda) - (ESPADA_TSet_PRF_A1 - (pair_EqDec - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (z <-$ (TSetCor_A F Enc A1); [t, l]<-2 z; ret (t, l, (t, l)))) - (ESPADA_TSet_PRF_A2 bigB bigS F_T - (fun (s : T lambda * list Blist) - (p : option (TSet) * list (Bvector lambda)) => - [t, q0]<-2 s; - [tSet_opt, tags]<-2 p; - tSet <- match tSet_opt with - | Some x => x - | None => nil - end; - ret (if tSet_opt then true else false) && - negb - (eqb (foreach (x in tags)TSetRetrieve tSet x) - (foreach (x in q0) - arrayLookupList (list_EqDec bool_EqDec) t x)))) - <= PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_2: - forall i, - PRF_NA_Advantage ({ 0 , 1 }^lambda) (RndF_R lambda bigB) F_T _ _ - (Hybrid.B1 nil _ _ (PRFI_A1 maxMatches (TSetCor_A F Enc A1)) i) - (Hybrid.B2 - (fun lsD => k <-$ {0, 1}^lambda; ret (map (F_T k) lsD)) - (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func (@RndF_range lambda bigB) _) nil lsD; ret lsR) _ - (PRFI_A2 lambda)) - <= PRF_NA_Adv. - - - Hypothesis maxKeywords_correct : - forall db q s, - In (db, q, s) (getSupport A1) -> - (length (ESPADA_SSE_SKS.toW db) <= maxKeywords)%nat. - - Hypothesis maxKeywords_correct_Cor : - forall (t : T lambda) (l : list Blist), - In (t, l) (getSupport (TSetCor_A F Enc A1)) -> - (length (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) <= - maxKeywords)%nat. - - Hypothesis maxMatches_correct_Cor : - forall (t : T lambda) (q : list Blist), - In (t, q) (getSupport (TSetCor_A F Enc A1)) -> - ESPADA_TSet_Correct_Once.maxMatch t maxMatches. - - Hypothesis bigB_correct : - forall (k : Bvector lambda) (b i : nat) (v1 : Bvector lambda) - (v2 : Bvector (S lambda)), (b, v1, v2) = F_T k i -> b < bigB. - - - Hypothesis failProb_correct_sec_A1: - forall (t : T lambda) (qs : list Blist) (s_A : (A_State * list (list (Bvector lambda)))), - In (t, qs, s_A) (getSupport (TSetSec_A1 F Enc A_State_EqDec A1)) -> - Pr[p <-$ ESPADA_TSetSetup_trial bigB bigS F_T F_bar_T t; ret !setupSuccess p ] <= - failProb. - - Hypothesis failProb_correct_sim_sec_A1 : - forall (t : T lambda) (qs : list Blist) - (s_A : A_State * list (list (Bvector lambda))), - In (t, qs, s_A) (getSupport (TSetSec_A1 F Enc A_State_EqDec A1)) -> - Pr - [p <-$ - ESPADA_TSet_Sim_trial lambda bigB bigS F_T - (ESPADA_TSet_Once_Games.L_T t qs) - (map (arrayLookupList (list_EqDec bool_EqDec) t) qs); - ret !simSetupSuccess p ] <= failProb. - - Hypothesis maxKeywords_correct_Sec : - forall (db : T lambda) (q : list Blist) - (s_A : A_State * list (list (Bvector lambda))), - In (db, q, s_A) (getSupport (TSetSec_A1 F Enc A_State_EqDec A1)) -> - (length (combineKeywords (removeDups (list_EqDec bool_EqDec) q) (toW db)) <= - maxKeywords)%nat. - - Hypothesis PRF_NA_Adv_correct_sec1 : - PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar_T - (list_EqDec bool_EqDec) (Bvector_EqDec lambda) - (Repeat_PRF_A1 - (pair_EqDec A_State_EqDec - (list_EqDec (list_EqDec (Bvector_EqDec lambda)))) - (TSetSec_A1 F Enc A_State_EqDec A1)) - (Repeat_PRF_A2 bigB bigS F_T F_bar_T (TSetSec_A2 A2) numTrials) <= - PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_sec2 : - forall i : nat, - PRF_NA_Advantage ({ 0 , 1 }^lambda) (RndF_range lambda bigB) F_T nat_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) - (B1 nil (list_EqDec nat_EqDec) - (pair_EqDec - (pair_EqDec - (pair_EqDec - (pair_EqDec - (pair_EqDec - (pair_EqDec - (pair_EqDec - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (pair_EqDec A_State_EqDec - (list_EqDec (list_EqDec (Bvector_EqDec lambda))))) - (list_EqDec (Bvector_EqDec lambda))) - (option_EqDec - (pair_EqDec - (list_EqDec - (list_EqDec - (option_EqDec - (pair_EqDec (Bvector_EqDec lambda) - (Bvector_EqDec (S lambda)))))) - (list_EqDec (list_EqDec nat_EqDec))))) - (list_EqDec (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (list_EqDec (list_EqDec bool_EqDec))) - (Repeat_IPRF_A1 - (pair_EqDec A_State_EqDec - (list_EqDec (list_EqDec (Bvector_EqDec lambda)))) bigB bigS F_T - (TSetSec_A1 F Enc A_State_EqDec A1)) i) - (B2 (fun lsD : list nat => k <-$ { 0 , 1 }^lambda; ret map (F_T k) lsD) - (fun lsD : list nat => - z <-$ - oracleMap - (list_EqDec - (pair_EqDec nat_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))))) - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) - (RndR_func (RndF_range lambda bigB) nat_EqDec) nil lsD; - [lsR, _]<-2 z; ret lsR) - (list_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda)))) - (Repeat_IPRF_A2 bigB bigS F_T F_bar_T (TSetSec_A2 A2) numTrials)) <= - PRF_NA_Adv. - - - Theorem ESPADA_SKS_Secure_concrete_T : - SSE_NA_Advantage SKS_EDBSetup_T _ SKS_Search_T A1 A2 (L Enc _ L_T) (SKS_Sim _ _ _ Sim_T) <= ((lambda / 1) * ((maxKeywords * (S maxMatches))^2 / 2 ^ lambda + - (maxKeywords / 1) * PRF_NA_Adv - + - PRF_NA_Adv) + (expRat failProb lambda)) + - (numTrials / 1 * - (PRF_NA_Adv + - (maxKeywords / 1) * - PRF_NA_Adv) + - expRat failProb numTrials + expRat failProb numTrials) + - PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ PRF_A_T + - (maxKeywords / 1) * IND_CPA_Adv. - - eapply leRat_trans. - eapply ESPADA_SKS_Secure_concrete; eauto. - eapply ratAdd_leRat_compat; intuition. - eapply ratAdd_leRat_compat; intuition. - eapply ratAdd_leRat_compat; intuition. - specialize ESPADA_TSet_Correct; intuition. - unfold TSetSetup, TSetGetTag, TSetRetrieve. - eapply leRat_trans. - Focus 2. - eapply H; - eauto. - - (* well-formedness *) - admit. - - reflexivity. - - specialize (@ESPADA_TSet_secure lambda (A_State * list (list (Bvector lambda))) _ bigB bigS F_T F_bar_T bigB_correct (TSetSec_A1 F Enc A_State_EqDec A1) (TSetSec_A2 A2)); intuition. - eapply leRat_trans. - eapply H; eauto. - - (* well-formedness*) - admit. - admit. - - reflexivity. - - reflexivity. - Qed. - - -End ESPADA_SSE_SKS_Secure_inst. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_SSE_SKS_Secure_auto.v b/src/ESPADA/ESPADA_SSE_SKS_Secure_auto.v deleted file mode 100644 index f3ef88b..0000000 --- a/src/ESPADA/ESPADA_SSE_SKS_Secure_auto.v +++ /dev/null @@ -1,738 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import SSE. -Require Import TSet. -Require Import RndPerm. -Require Import CompFold. -Require Import Array. -Require Import PRF. -Require Import Encryption. -Require Import OracleCompFold. -Require Import ESPADA_SSE_SKS. -Require Import ESPADA_SSE_SKS_Secure_tacs. - -Local Open Scope list_scope. - - - - -Section ESPADA_SSE_SKS_Secure. - - Variable lambda : posnat. - - (* PRF *) - Variable F : Bvector lambda -> Keyword -> Bvector lambda. - - (* Encryption scheme *) - Variable Enc : Bvector lambda -> Bvector lambda -> Comp (Bvector lambda). - Variable Dec : Bvector lambda -> Bvector lambda -> Bvector lambda. - - (* TSet *) - Variable TSet : Set. - Hypothesis TSet_EqDec : EqDec TSet. - Variable TSetKey : Set. - Hypothesis TSetKey_EqDec : EqDec TSetKey. - Variable Tag : Set. - Hypothesis Tag_EqDec : EqDec Tag. - Variable TSetSetup : T (pos (lambda)) -> Comp (TSet * TSetKey). - Variable TSetGetTag : TSetKey -> Keyword -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector (lambda))). - Variable Leakage_T : Set. - Hypothesis Leakage_T_EqDec : EqDec Leakage_T. - Variable L_T : T (pos (lambda)) -> list Keyword -> Leakage_T. - Variable Sim_T : Leakage_T -> list (list (Bvector lambda)) -> Comp (TSet * list Tag). - - Definition EDB := EDB TSet. - Definition SearchTranscript := SearchTranscript lambda Tag. - Definition SKS_EDBSetup := SKS_EDBSetup F Enc _ _ TSetSetup. - Definition SKS_Search := @SKS_Search lambda F Dec TSet TSetKey Tag _ - TSetGetTag TSetRetrieve. - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (DB lambda * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall x y z, well_formed_comp (A2 x y z). - - Notation "'gInst' g" := (g lambda F Enc TSet TSetKey Tag _ TSetSetup TSetGetTag TSetRetrieve _ A1 A2) (at level 90). - - (* Step 1: inline and simplify *) - Theorem Real_G1_equiv : - Pr[SSE_Sec_NA_Real SKS_EDBSetup _ SKS_Search A1 A2] == Pr[gInst G1]. - - unfold G1, SSE_Sec_NA_Real, SKS_EDBSetup, SKS_Search, SKS_Search_1, ESPADA_SSE_SKS.SKS_Search. - do 4 skip. - - eapply comp_spec_eq_impl_eq. - - repeat specTac. - - Qed. - - (* Step 2: factor out some TSet computations and prepare to apply TSet correctness*) - - Definition SKS_EDBSetup_wLoop := SKS_EDBSetup_wLoop F Enc. - - Definition G1_1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - ls <-$ ( - tags <-$ compMap _ (TSetGetTag k_T) q; - ret (map (fun z => (z, TSetRetrieve tSet z)) tags) - ); - A2 s_A tSet ls. - - - Theorem G1_G1_1_equiv : - Pr[gInst G1] == Pr[G1_1]. - - unfold G1, G1_1, SKS_EDBSetup_wLoop. - do 4 skip. - comp_skip. - - symmetry. - eapply comp_spec_eq_impl_eq. - - eapply compMap_map_fission_eq; intuition. - repeat specTac. - - Qed. - - Theorem G1_1_G2_equiv : - Pr[G1_1] == Pr[gInst G2]. - - unfold G1_1, G2, SKS_EDBSetup_wLoop. - do 5 skip. - - eapply eqRat_refl_eq. - f_equal. - f_equal. - - repeat lpTac. - - Qed. - - Theorem G1_G2_equiv : - Pr[gInst G1] == Pr[gInst G2]. - - rewrite G1_G1_1_equiv. - eapply G1_1_G2_equiv. - - Qed. - - (* Step 3: apply TSet correctness. *) - Definition G2_1 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - v <-$ A2 s_A tSet ls; - ret (v, negb (eqb rs rs')). - - Definition G2_2 := - [db, q, s_A] <-$3 A1; - k_S <-$ {0, 1}^lambda; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ls <- combine tags rs'; - v <-$ A2 s_A tSet ls; - ret (v, negb (eqb rs rs')). - - Theorem G2_G2_1_equiv : - Pr[gInst G2] == Pr[x <-$ G2_1; ret fst x]. - - unfold G2, G2_1, SKS_EDBSetup_wLoop. - repeat skip. - inline_first. - - identSkip_l; comp_simp; simpl; intuition. - - Qed. - - Theorem combine_eq_parts : - forall (A B : Set)(a1 a2 : list A)(b1 b2 : list B), - a1 = a2 -> - b1 = b2 -> - combine a1 b1 = combine a2 b2. - - intuition. - subst. - trivial. - - Qed. - - Theorem G2_1_G2_2_eq_until_bad : - forall x, - evalDist G2_1 (x, false) == evalDist G2_2 (x, false). - - intuition. - unfold G2_1, G2_2. - do 5 skip. - - repeat distTac; eauto using combine_eq_parts. - - Qed. - - Theorem G2_1_G2_2_bad_same : - Pr[x <-$ G2_1; ret snd x] == Pr[x <-$ G2_2; ret snd x]. - - unfold G2_1, G2_2. - do 5 skip. - inline_first. - comp_irr_l. comp_irr_r. - comp_simp; simpl; intuition. - - Qed. - - Definition G2_1_bad := - [t, q] <-$2 (TSetCor_A F Enc A1); - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - rs <- map (TSetRetrieve tSet) tags; - rs' <- map (arrayLookupList _ t) q; - ret (negb (eqb rs rs')). - - Theorem G2_1_bad_equiv : - Pr[x <-$ G2_1; ret snd x] == Pr[G2_1_bad]. - - unfold G2_1, G2_1_bad, TSetCor_A. - repeat skip. - reflexivity. - inline_first. - comp_irr_l. - simpl; intuition. - - Qed. - - Theorem G2_1_bad_small : - Pr[G2_1_bad] <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - eapply leRat_refl. - - Qed. - - Theorem G2_1_G2_2_close : - | Pr[x <-$ G2_1; ret fst x] - Pr[x <-$ G2_2; ret fst x] | <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - intuition. - - eauto using leRat_trans, fundamental_lemma_h, eqRat_impl_leRat, - G2_1_bad_equiv, G2_1_bad_small, G2_1_G2_2_eq_until_bad, G2_1_G2_2_bad_same. - - Qed. - - Notation "'gInst1' g" := (g lambda F Enc TSet TSetKey Tag _ TSetSetup TSetGetTag _ A1 A2) (at level 90). - - Theorem G2_2_G3_equiv : - Pr[x <-$ G2_2; ret fst x] == Pr[gInst1 G3]. - - unfold G2_2, G3. - repeat skip. - reflexivity. - inline_first. - - identSkip_r. - comp_simp; simpl; intuition. - - Qed. - - Theorem G2_G3_equiv : - | Pr[gInst G2] - Pr[gInst1 G3] | <= AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1). - - rewrite G2_G2_1_equiv. - rewrite <- G2_2_G3_equiv. - eapply G2_1_G2_2_close. - - Qed. - - - (* Step 4: apply TSet security. *) - Definition TSetSec_A1 := TSetSec_A1 F Enc _ A1. - Definition TSetSec_A2 := TSetSec_A2 A2. - Definition G3_1 := - [t, q, s] <-$3 TSetSec_A1; - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - TSetSec_A2 s (tSet, tags). - - Definition G3_2 := - [t, q, s] <-$3 TSetSec_A1; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - TSetSec_A2 s (tSet, tags). - - Theorem G3_G3_1_equiv : - Pr[gInst1 G3] == Pr[G3_1]. - - unfold G3, G3_1, TSetSec_A1, TSetSec_A2. - repeat skip. - reflexivity. - - Qed. - - Theorem G3_1_G3_2_close: - | Pr[G3_1] - Pr[G3_2] | <= TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T. - - eapply leRat_refl. - - Qed. - - Notation "'gInst2' g" := (g lambda F Enc TSet Tag _ L_T Sim_T _ A1 A2) (at level 90). - - Theorem G3_2_G4_equiv : - Pr[G3_2] == Pr[gInst2 G4]. - - unfold G3_2, G4, TSetSec_A1, TSetSec_A2. - repeat (inline_first; comp_skip; comp_simp). - reflexivity. - Qed. - - Theorem G3_G4_equiv : - | Pr[gInst1 G3] - Pr[gInst2 G4] | <= TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T. - - rewrite G3_G3_1_equiv. - rewrite <- G3_2_G4_equiv. - eapply G3_1_G3_2_close. - - Qed. - - (* Step 5-6: apply PRF definition. *) - Definition G4_1 := - k_S <-$ {0, 1}^lambda; - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop db k_S) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Definition SKS_EDBSetup_wLoop_5 := SKS_EDBSetup_wLoop_5 Enc. - - Theorem SKS_EDBSetup_5_spec : - forall d b z x, - comp_spec - (fun (x0 : Keyword * list (Bvector lambda)) - (y : Keyword * list (Bvector lambda) * unit) => - x0 = (fst y)) (SKS_EDBSetup_wLoop d x b) - ((SKS_EDBSetup_wLoop_5 d b) unit unit_EqDec - (f_oracle F (Bvector_EqDec lambda) x) z). - - intuition. - unfold SKS_EDBSetup_wLoop, SKS_EDBSetup_wLoop_5, f_oracle. - simpl. comp_simp. - skip. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G4_G4_1_equiv : - Pr[gInst2 G4] == Pr[G4_1]. - - unfold G4, G4_1. - comp_swap_r. - repeat (comp_skip; comp_simp). - reflexivity. - Qed. - - Theorem G4_1_G5_equiv : - Pr[G4_1] == Pr[gInst2 G5]. - - Local Opaque evalDist. - - unfold G4_1, G5, PRF_A. - repeat (inline_first; comp_skip; comp_simp; simpl). - inline_first. - eapply comp_spec_eq_impl_eq. - comp_skip; [specTac | idtac]. - eapply SKS_EDBSetup_5_spec. - comp_simp. - lpTac. - skip. - simpl. inline_first. - identSkip_l. comp_simp. specTac. - - Qed. - - Theorem G4_G5_equiv : - Pr[gInst2 G4] == Pr[gInst2 G5]. - - rewrite G4_G4_1_equiv. - eapply G4_1_G5_equiv. - - Qed. - - Notation "'gInst3' g" := (g lambda Enc TSet Tag _ L_T Sim_T _ A1 A2) (at level 90). - Definition PRF_A := PRF_A Enc L_T Sim_T A1 A2. - - Theorem G5_G6_equiv : - | Pr[gInst2 G5] - Pr[gInst3 G6] | <= PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ PRF_A. - - eapply leRat_refl. - - Qed. - - Definition SKS_EDBSetup_wLoop_6_f db w k_e := - inds <- lookupInds db w; - t <-$ compMap _ (SKS_EDBSetup_indLoop _ Enc k_e) inds; - ret (w, t). - - Definition SKS_EDBSetup_wLoop_6 db w := - k_e <--$ query w; - $ (SKS_EDBSetup_wLoop_6_f db w k_e). - - Definition PRF_A_6 := - [db, q, s_A] <--$3$ A1; - t <--$ oc_compMap _ (SKS_EDBSetup_wLoop_6 db) (toW db); - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <--$2$ Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - $ A2 s_A tSet ls. - - Definition G6_1 := - [b, _] <-$2 PRF_A_6 _ _ (RndR_func (Rnd lambda) _) nil; - ret b. - - - Theorem G6_G6_1_equiv : - Pr[gInst3 G6] == Pr[G6_1]. - - unfold G6, G6_1. - - unfold PRF_A, PRF_A_6, SKS_EDBSetup_wLoop_5, SKS_EDBSetup_wLoop_6, SKS_EDBSetup_wLoop_6_f. - do 2 (simpl; inline_first; comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - specTac; simpl. - repeat skip. - eapply comp_spec_ret; intuition. - Qed. - - Definition G6_2 := - [b, _] <-$2 PRF_A_6 _ _ (fun s a => x <-$ Rnd lambda; ret (x, s)) tt; - ret b. - - Theorem G6_1_G6_2_equiv : - Pr[G6_1] == Pr[G6_2]. - - unfold G6_1, G6_2, PRF_A_6. - simpl; skip. - simpl; inline_first. - eapply comp_spec_eq_impl_eq. - specTac. - eapply compMap_randomFunc_NoDup; intuition; eapply removeDups_NoDup. - - inline_first. comp_simp. inline_first. simpl in *. subst. - repeat (skip; simpl). - eapply comp_spec_eq_refl. - - Qed. - - - (* Step 7: replace RF outputs with random values. *) - - Theorem EDBSetup_7_spec : - forall d b z, - comp_spec - (fun (x : Keyword * list (Bvector lambda)) - (y : Keyword * list (Bvector lambda) * unit) => - x = (fst y)) (SKS_EDBSetup_wLoop_7 Enc d b) - ((SKS_EDBSetup_wLoop_6 d b) unit unit_EqDec - (fun (s : unit) (_ : Keyword) => x <-$ { 0 , 1 }^lambda; ret (x, s)) - z). - - intuition. - unfold SKS_EDBSetup_wLoop_6, SKS_EDBSetup_wLoop_7, SKS_EDBSetup_wLoop_6_f. - repeat (simpl; skip). - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G6_2_G7_equiv : - Pr[G6_2] == Pr[gInst3 G7]. - - unfold G6_2, G7, PRF_A_6. - eapply comp_spec_eq_impl_eq. - do 2 (simpl; skip). - specTac. - eapply EDBSetup_7_spec. - inline_first; simpl in *; lpTac. - skip. simpl. inline_first. - - identSkip_r. - - Qed. - - Theorem G6_G7_equiv : - Pr[gInst3 G6] == Pr[gInst3 G7]. - - rewrite G6_G6_1_equiv. - rewrite G6_1_G6_2_equiv. - eapply G6_2_G7_equiv. - - Qed. - - (* Step 8: apply encryption security definition. *) - Definition SKS_EDBSetup_wLoop_7_1 db w := - k_e <-$ (Rnd lambda); - inds <- lookupInds db w; - compMap _ (SKS_EDBSetup_indLoop _ Enc k_e) inds. - - Definition G7_1 := - [db, q, s_A] <-$3 A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_1 db) (toW db); - t <- combine (toW db) t; - inds <- map (lookupAnswers _ t) q; - [tSet, tags] <-$2 Sim_T (L_T t q) inds; - rs <- map (arrayLookupList _ t) q; - ls <- combine tags rs; - A2 s_A tSet ls. - - Theorem EDBSetup_7_1_spec : - forall d b, - comp_spec (fun a b => snd a = b) (SKS_EDBSetup_wLoop_7 Enc d b) (SKS_EDBSetup_wLoop_7_1 d b). - - intuition. - unfold SKS_EDBSetup_wLoop_7, SKS_EDBSetup_wLoop_7_1. - skip. - identSkip_r. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G7_G7_1_equiv : - Pr[gInst3 G7] == Pr[G7_1]. - - unfold G7, G7_1, SKS_EDBSetup_wLoop_7. - eapply comp_spec_eq_impl_eq. - do 2 skip. - specTac. - eapply EDBSetup_7_1_spec. - lpTac. - eapply (@compMap_support _ _ (fun a b => fst b = a)) in H1. - lpTac. - specialize (split_combine a0); intuition. - remember (split a0) as z. destruct z. simpl in *. subst. - specTac. - - intuition. repeat simp_in_support. trivial. - - Qed. - - - Definition SKS_EDBSetup_wLoop_7_2 w := - k_e <-$ (Rnd lambda); - [b, _] <-$2 (Enc_A2 w) _ _ (EncryptOracle Enc _ k_e) tt; - ret b. - - Definition G7_2 := - [lsa, s_A] <-$2 (Enc_A1 _ A1); - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_2) lsa; - Enc_A3 L_T Sim_T A2 s_A t. - - Theorem EncOracle_spec : - forall b0 z x1, - comp_spec - (fun (x2 : Bvector lambda) (y : Bvector lambda * unit) => - x2 = (fst y)) (Enc x1 b0) - ((query b0) unit unit_EqDec - (EncryptOracle Enc (Bvector_EqDec lambda) x1) z). - - intuition. - unfold EncryptOracle. - simpl. - identSkip_l. - eapply comp_spec_ret; intuition. - - Qed. - - Theorem G7_1_G7_2_equiv : - Pr[G7_1] == Pr[G7_2]. - - unfold G7_1, G7_2, Enc_A1, SKS_EDBSetup_wLoop_7_1, SKS_EDBSetup_wLoop_7_2, Enc_A2, Enc_A3. - do 2 skip. - distTac. repeat lpTac. intuition. - destruct H0. intuition. subst. - distTac. - eapply comp_spec_eq_impl_eq. - identSkip_l. - specTac. - eapply EncOracle_spec. - comp_simp. lpTac. - eapply comp_spec_eq_refl. - reflexivity. - Qed. - - - Definition SKS_EDBSetup_wLoop_7_3 w := - k_e <-$ (Rnd lambda); - [b, _] <-$2 (Enc_A2 w) _ _ (EncryptNothingOracle Enc _ (zeroVector lambda) k_e) tt; - ret b. - - Definition G7_3 := - [lsa, s_A] <-$2 Enc_A1 _ A1; - t <-$ compMap _ (SKS_EDBSetup_wLoop_7_3) lsa; - Enc_A3 L_T Sim_T A2 s_A t. - - Theorem G7_2_G7_3_close : - | Pr[G7_2] - Pr[G7_3] | <= - IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc (Enc_A1 _ A1) (@Enc_A2 lambda) (Enc_A3 L_T Sim_T A2) _ _(zeroVector lambda). - - eapply leRat_refl. - Qed. - - Theorem EncNothing_spec : - forall b0 z x1, - comp_spec - (fun (x2 : Bvector lambda) (y : Bvector lambda * unit) => - x2 = (fst y)) (Enc x1 (zeroVector lambda)) - ((query b0) unit unit_EqDec - (EncryptNothingOracle Enc (Bvector_EqDec lambda) - (zeroVector lambda) x1) z). - - intuition. - unfold EncryptNothingOracle. - simpl. identSkip_l. - eapply comp_spec_ret; intuition. - Qed. - - Notation "'gInst4' g" := (g _ Enc _ _ _ L_T Sim_T _ A1 A2) (at level 90). - - Theorem G7_3_G8_equiv : - Pr[G7_3] == Pr[gInst4 G8]. - - unfold G7_3, G8, Enc_A1, SKS_EDBSetup_wLoop_7_3, SKS_resultsStruct, Enc_A2, Enc_A3. - do 2 skip. distTac. repeat lpTac. intuition. - destruct H0; intuition; pairInv. - skip. - eapply comp_spec_eq_impl_eq. - identSkip_r. - specTac. - eapply EncNothing_spec. - simpl in *. lpTac. - reflexivity. - - Qed. - - Theorem G7_G8_equiv : - | Pr[gInst3 G7] - Pr[gInst4 G8] | <= - IND_CPA_SecretKey_OI_Advantage (Rnd lambda) Enc (Enc_A1 _ A1) (@Enc_A2 lambda) (Enc_A3 L_T Sim_T A2) _ _(zeroVector lambda). - - rewrite G7_G7_1_equiv. - rewrite G7_1_G7_2_equiv. - rewrite <- G7_3_G8_equiv. - eapply G7_2_G7_3_close. - - Qed. - - - Theorem G8_Ideal_equiv : - Pr[gInst4 G8] == Pr[SSE_Sec_NA_Ideal A1 A2 (L Enc _ L_T) (SKS_Sim _ _ _ Sim_T)]. - - Local Opaque evalDist. - unfold G8, SSE_Sec_NA_Ideal, L, SKS_Sim. - do 3 skip; - reflexivity. - - Qed. - - Variable IND_CPA_Adv : Rat. - - Hypothesis IND_CPA_Adv_correct : - (forall i, - IND_CPA_SecretKey_O_Advantage ({0,1}^lambda) Enc - (A_c - (B1 (nil, nil) _ _ (Enc_A1 _ A1) i) - (@Enc_A2 lambda) - (B2 - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) _ _ (EncryptOracle Enc _ key) tt; - [b, _]<-2 z; ret b) - (fun x => - key <-$ {0,1}^lambda; - z <-$ - (Enc_A2 x) unit unit_EqDec - (EncryptNothingOracle Enc _ (zeroVector lambda) key) tt; - [b, _]<-2 z; ret b) _ (Enc_A3 L_T Sim_T A2)) - ) - _ - (zeroVector lambda) - <= IND_CPA_Adv)%rat. - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - forall db q s, - In (db, q, s) (getSupport A1) -> - (length (toW db) <= maxKeywords)%nat. - - Hypothesis Enc_wf : - forall k p, - well_formed_comp (Enc k p). - - Theorem ESPADA_SKS_Secure_concrete : - SSE_NA_Advantage SKS_EDBSetup _ SKS_Search A1 A2 (L Enc _ L_T) (SKS_Sim _ _ _ Sim_T) <= - AdvCor _ TSetSetup TSetGetTag TSetRetrieve (TSetCor_A F Enc A1) + - TSetAdvantage _ TSetSetup TSetGetTag L_T TSetSec_A1 TSetSec_A2 Sim_T + - PRF_Advantage (Rnd lambda) (Rnd lambda) F _ _ PRF_A + - (maxKeywords / 1) * IND_CPA_Adv. - - unfold SSE_NA_Advantage. - rewrite Real_G1_equiv. - rewrite G1_G2_equiv. - eapply leRat_trans. - eapply ratDistance_le_trans. - eapply G2_G3_equiv. - eapply ratDistance_le_trans. - eapply G3_G4_equiv. - rewrite G4_G5_equiv. - eapply ratDistance_le_trans. - eapply G5_G6_equiv. - rewrite G6_G7_equiv. - rewrite <- G8_Ideal_equiv. - eapply G7_G8_equiv. - repeat rewrite <- ratAdd_assoc. - - eapply ratAdd_leRat_compat. intuition. - Ltac oc_wftac := - match goal with - | [|- well_formed_oc _ ] => eapply oc_compMap_wf; intuition - | [|- well_formed_oc (OC_Query _ _)] => econstructor - end. - eapply IND_CPA_OI_impl_O. intuition. wftac. intuition. - repeat oc_wftac. - intuition. - unfold Enc_A1 in *. - repeat simp_in_support. - destruct x. destruct p. - repeat simp_in_support. - rewrite map_length. - eauto. - - eauto. - - intuition. - - Qed. - - Print Assumptions ESPADA_SKS_Secure_concrete. - -End ESPADA_SSE_SKS_Secure. - - diff --git a/src/ESPADA/ESPADA_SSE_SKS_Secure_tacs.v b/src/ESPADA/ESPADA_SSE_SKS_Secure_tacs.v deleted file mode 100644 index 676ed16..0000000 --- a/src/ESPADA/ESPADA_SSE_SKS_Secure_tacs.v +++ /dev/null @@ -1,71 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import CompFold. -Require Import OracleCompFold. - -Ltac skip := (inline_first; comp_skip; comp_simp). - -Ltac identSkipSpec_l := eapply comp_spec_eq_trans_l; [eapply comp_spec_eq_symm; eapply comp_spec_right_ident | idtac]; comp_skip. - -Ltac identSkipSpec_r := eapply comp_spec_eq_trans_r; [idtac | eapply comp_spec_right_ident]; comp_skip. - -Ltac identSkip_l := - match goal with - | [|- evalDist ?x _ == evalDist _ _] => rewrite <- (@evalDist_right_ident _ _ x); comp_skip - | [|- comp_spec _ _ _ ] => identSkipSpec_l - end. - -Ltac identSkip_r := - match goal with - | [|- evalDist _ _ == evalDist ?x _] => rewrite <- (@evalDist_right_ident _ _ x); comp_skip - | [|- comp_spec _ _ _ ] => identSkipSpec_r - end. - -Ltac lpTac := - try match goal with - | [|- list_pred _ ?a ?a] => eapply list_pred_eq - | [H : list_pred (fun x y => snd x = y) _ _ |- _ ] => apply snd_split_eq_list_pred in H - | [H : list_pred (fun x y => fst y = x) _ _ |- _ ] => apply list_pred_symm in H; apply fst_split_eq_list_pred in H - | [ |- map _ ?ls = combine ?ls (map _ ?ls) ] => rewrite combine_map_eq - | [ |- map _ _ = map _ _ ] => eapply map_ext_pred - | [ |- _] => eapply list_pred_combine_r - | [ H : fst ?b = snd ?b |- _ ] => destruct b; simpl in *; subst - | [ H : list_pred eq _ _ |- _] => eapply list_pred_eq_impl_eq in H; subst - | [ |- list_pred _ _ (map _ _ )] => eapply list_pred_map_r - | [ |- list_pred _ (map _ _ ) _] => eapply list_pred_map_l - end; intuition; subst. - -Ltac existTac := - match goal with - | [H : Comp ?A |- ?A ] => eapply comp_base_exists; intuition - end. - -Ltac specTac := - try match goal with - | [ |- comp_spec _ _ _ ] => eapply comp_spec_eq_refl - | [ |- comp_spec _ (Bind _ _) (Bind _ _)] => comp_skip; try existTac - | [ |- comp_spec _ (compMap _ _ _ ) (compMap _ _ _ ) ] => eapply compMap_spec - | [ |- comp_spec _ (Ret _ (?b, ?a)) (Ret _ ?a )] => eapply (comp_spec_ret _ _ (fun a b => snd a = b)); intuition - | [|- comp_spec _ ?a ?a] => eapply comp_spec_consequence; [eapply comp_spec_eq_refl | intuition] - | [|- comp_spec eq ((evalDist_OC (oc_compMap _ _ _)) _ _ _ _) ((evalDist_OC (oc_compMap _ _ _)) _ _ _ _)] => eapply oc_compMap_eq - | [|- comp_spec _ (compMap _ _ _) ((evalDist_OC (oc_compMap _ _ _)) _ _ _ _)] => eapply compMap_oc_spec - | [|- comp_spec _ ((evalDist_OC (oc_compMap _ _ _)) _ _ _ _) (compMap _ _ _)] => eapply comp_spec_symm; eapply compMap_oc_spec - | [H : eqb _ _ = true |- _] => rewrite eqb_leibniz in H - | [|- comp_spec _ (?f ?a) (?f ?b) ] => case_eq (eqb a b); intuition - end; try lpTac; subst. - -Ltac distTac := - match goal with - | [H : eqb _ _ = true |- _] => rewrite eqb_leibniz in H - | [H : negb _ = false |- _] => apply negb_false_iff in H - | [ |- evalDist (Bind ?a _) _ == evalDist (Bind ?a _) _] => comp_skip - | [H : ?a = ?b |- evalDist (Bind (?f ?a) _) _ == evalDist (Bind (?f ?b) _ ) _ ] => rewrite H - | [H : ?a <> ?b |- evalDist (Bind (?f ?a) _) _ == evalDist (Bind (?f ?b) _ ) _ ] => comp_irr_l; comp_irr_r; repeat rewrite evalDist_ret_0; intuition; pairInv - | [|- evalDist (Bind (?f ?a) _) _ == evalDist (Bind (?f ?b) _ ) _] => destruct (EqDec_dec _ a b) - | [ |- evalDist (compMap _ _ _) _ == evalDist (compMap _ _ _) _ ] => eapply compMap_eq - end. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet.v b/src/ESPADA/ESPADA_TSet.v deleted file mode 100644 index c275341..0000000 --- a/src/ESPADA/ESPADA_TSet.v +++ /dev/null @@ -1,148 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import FCF. -Require Import Array. -Require Import CompFold. -Require Import RndNat. -Require Import RndListElem. - -Require Import TSet. -Require Import PRF. - -Local Open Scope array_scope. -Local Open Scope list_scope. - -Notation "! x" := (negb x) (at level 50). -Infix "=?" := eqb (at level 70) : comp_scope. -Notation " x '!=' y " := (negb (eqb x y)) (at level 70) : comp_scope. - -Notation "'loop_over' ( init , ls ) c " := - (compFold _ (foldBody_option _ c%comp) (Some init) ls) - (right associativity, at level 85) : comp_scope. - Notation "'loop_over*' ( init , ls ) c " := - (compFold _ (foldBody_option _ c%comp) (init) ls) - (right associativity, at level 85) : comp_scope. - - -(* The ESPADA TSet implementation. *) -Section ESPADA_TSet. - - (* For simplicity, we use lambda as the bit length of the strings in T. We could generalize this to a polynomial function of lambda. *) - Variable lambda : posnat. - - (* A TSet is an array of B buckets of size N. The protocol determines B and S from the size of the data structure that will be encrypted. We make B and S variables because several things depend on them. We prove security for all B S, and for all adversaries that provide a structure of size N such that B and S are correct with respect to N. *) - Variable tSize bigB bigS : posnat. - Variable bucketSize : nat -> (posnat * posnat). - - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - (* F maps a pair of (tag, nat) to a nat (indicating the bucket) and a couple of bit vectors. The bucket index must be less than the number of buckets. We could enforce this using a sigma type, but that doesn't match the definition for RndNat. *) - Hypothesis F_b_correct : - forall k b i v1 v2, - (b, v1, v2) = F k i -> - b < bigB. - - - Definition TSet_Record := (Bvector lambda * Bvector (S lambda))%type. - Definition TSet := list (list (option TSet_Record)). - - (* The TSet is constructed using a helper structure that keeps track of which positions are free in the TSet. *) - Definition FreeList := list (list nat). - - (* toW returns all the keywords in the structure. We removeDups because an array may have additional (dead) entries for some keyword. *) - Definition toW (bigT : T lambda) := - removeDups _ (fst (unzip bigT)). - - (* bigN is the total size of the structure. Note that we can't simply sum over all entries in the structure because there may be dead entries. *) - Definition bigN (bigT : T lambda) := - bigW <- toW bigT; - fold_left (fun acc w => acc + length (arrayLookupList _ bigT w))%nat bigW O. - - (* tSetUpdate replace the value at some location in the TSet. *) - Definition tSetUpdate (tSet : TSet)(b j : nat)(r : TSet_Record) : TSet := - let bucket := nth b tSet nil in - let bucket' := listReplace bucket j (Some r) None in - listReplace tSet b bucket' nil. - - - Local Open Scope vector_scope. - Definition ESPADA_TSetSetup_tLoop (stag : Bvector lambda)(length : nat)(acc : TSet * FreeList)(e : nat * Bvector lambda) := - [tSet, free] <-2 acc; - [i, s_i] <-2 e; - [b, bigL, bigK] <-3 F stag i; - free_b <- nth b free nil; - j <-$? (rndListElem _ free_b) ; - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - bet <- (S i) != length; - newRecord <- (bigL, (bet :: s_i) xor bigK); - tSet <- tSetUpdate tSet b j newRecord; - ret (tSet, free). - - Definition initFree := - for ( _ '< bigB) - (for (i '< bigS) i). - - Definition ESPADA_TSetSetup_wLoop (bigT : T lambda) (k_T : Bvector lambda)(acc : TSet * FreeList)(w : Blist) := - [tSet, free] <-2 acc; - stag <- F_bar k_T w; - t <- lookupAnswers _ bigT w; - ls <- zip (allNatsLt (length t)) t; - loop_over ((tSet, free), ls) (ESPADA_TSetSetup_tLoop stag (length t)). - - - Definition ESPADA_TSetSetup_trial (bigT : T lambda) := - k_T <-$ {0, 1} ^ lambda; - loopRes <-$ loop_over ((nil, initFree), (toW bigT)) (ESPADA_TSetSetup_wLoop bigT k_T) ; - ret (loopRes, k_T). - - Definition getTSet(opt : option (TSet * FreeList)) := - match opt with - | None => nil - | Some p => fst p - end. - - Definition isSome (A : Set)(opt : option A) := - if opt then true else false. - - Definition ESPADA_TSetSetup t := - [res, k_T] <-$2 Repeat (ESPADA_TSetSetup_trial t) (fun p => isSome (fst p)); - ret (getTSet res, k_T). - - Definition ESPADA_TSetGetTag (k_T : Bvector lambda) w := - ret (F_bar k_T w). - - Definition ESPADA_TSetRetrieve_tLoop tSet stag i := - [b, L, K] <-3 F stag i; - B <- nth b tSet nil; - t <- arrayLookupOpt _ B L; - match t with - | None => None - | Some u => - v <- u xor K; - bet <- Vector.hd v; - s <- Vector.tl v; - Some (s, bet) - end. - - Local Open Scope list_scope. - Fixpoint ESPADA_TSetRetrieve_h tSet stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop tSet stag i) with - | Some (v, bet) => - v :: (if (bet) then (ESPADA_TSetRetrieve_h tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Variable maxMatches : nat. - - Definition ESPADA_TSetRetrieve tSet stag := - ESPADA_TSetRetrieve_h tSet stag O maxMatches. - -End ESPADA_TSet. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Correct.v b/src/ESPADA/ESPADA_TSet_Correct.v deleted file mode 100644 index 08a1f4e..0000000 --- a/src/ESPADA/ESPADA_TSet_Correct.v +++ /dev/null @@ -1,1002 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import ESPADA_TSet_Correct_Once_Games. -Require Import ESPADA_TSet_Correct_Once. -Require Import TSet. -Require Import ESPADA_TSet_Once_Games. -Require Import CompFold. - - -Section ESPADA_TSet_Correct. - - Variable lambda : posnat. - Variable bigB bigS maxMatches : posnat. - - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Variable A : Comp (T lambda * list Blist). - - Definition ESPADA_TSetSetup_trial := ESPADA_TSetSetup_trial bigB bigS F F_bar. - Definition ESPADA_TSetRetrieve := ESPADA_TSetRetrieve lambda F maxMatches. - - Variable failProb : Rat. - Hypothesis failProb_correct : - forall t l, - In (t, l) (getSupport A) -> - Pr[x <-$ ESPADA_TSetSetup_trial t; ret (if (fst x) then false else true)] <= failProb. - Hypothesis failProb_lt_1 : - ~ (1 <= failProb). - - Hypothesis bigB_correct : - forall (k : Bvector lambda) (b i : nat) (v1 : Bvector lambda) - (v2 : Bvector (S lambda)), (b, v1, v2) = F k i -> b < bigB. - - Require Import ESPADA_TSet_Secure. - - - Definition TSetCorr_G0_RepeatBody z := - ([res, k_T] <-$2 ESPADA_TSetSetup_trial (fst z); - tags <-$ (foreach (x in (snd z))ESPADA_TSetGetTag lambda F_bar k_T x); - tSet <- getTSet res; - t_w <- (foreach (x in tags)ESPADA_TSetRetrieve tSet x); - - t_w_correct <- - (foreach (x in (snd z))Array.arrayLookupList (list_EqDec bool_EqDec) (fst z) x); - b <- negb (eqb t_w t_w_correct); - ret (res, (if res then true else false) && b)). - - Definition TSetCorr_G1 := - z <-$ A; - z0 <-$ - Repeat - (TSetCorr_G0_RepeatBody z) - (fun p => isSome (fst p)); - ret ( snd z0). - - Theorem ESPADA_TSet_Correct_G1_equiv : - AdvCor _ - (ESPADA_TSetSetup bigB bigS F F_bar) - (ESPADA_TSetGetTag lambda F_bar) - (ESPADA_TSetRetrieve) - A == - Pr[TSetCorr_G1]. - - unfold AdvCor, AdvCor_G, TSetCorr_G1. - comp_skip. - comp_simp. - unfold ESPADA_TSetSetup. - inline_first. - - Theorem Repeat_factor_indep : - forall (A B : Set)(eqda : EqDec A)(eqdb : EqDec B)(c1 : Comp A)(c2 : A -> Comp B) P x, - well_formed_comp c1 -> - (exists a, In a (filter P (getSupport c1))) -> - (forall a, In a (getSupport c1) -> well_formed_comp (c2 a)) -> - evalDist (Repeat (a <-$ c1; b <-$ (c2 a); ret (a, b)) (fun p => P (fst p))) x == - evalDist (a <-$ Repeat c1 P; b <-$ (c2 a); ret (a, b)) x. - - intuition. - symmetry. - eapply repeat_fission'; intuition. - wftac. - repeat simp_in_support. - trivial. - - Qed. - - unfold TSetCorr_G0_RepeatBody. - - assert ( - Pr - [z0 <-$ - Repeat - (z0 <-$ ESPADA_TSetSetup_trial (fst (t, l)); - [res, k_T]<-2 z0; - tags <-$ - compMap (Bvector_EqDec lambda) - (fun x : Blist => ESPADA_TSetGetTag lambda F_bar k_T x) - (snd (t, l)); - tSet <- getTSet res; - t_w <- - map (fun x : Bvector lambda => ESPADA_TSetRetrieve tSet x) tags; - t_w_correct <- - map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) (fst (t, l)) x) - (snd (t, l)); - b <- negb (eqb t_w t_w_correct); - ret (res, (if res then true else false) && b)) - (fun p : option (TSet lambda * FreeList) * bool => isSome (fst p)); - ret snd z0 ] - == - Pr - [z0 <-$ - Repeat - (z0 <-$ ESPADA_TSetSetup_trial t; - z1 <-$ ( - tags <-$ - compMap (Bvector_EqDec lambda) - (fun x : Blist => ESPADA_TSetGetTag lambda F_bar (snd z0) x) - l; - tSet <- getTSet (fst z0); - t_w <- - map (fun x : Bvector lambda => ESPADA_TSetRetrieve tSet x) tags; - t_w_correct <- - map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) - l; - b <- negb (eqb t_w t_w_correct); - ret ((if (fst z0) then true else false) && b) - ); - ret (z0, z1)) - (fun p => isSome (fst (fst p))); - ret snd z0 ] - ). - - eapply comp_spec_eq_impl_eq. - comp_skip. - - Theorem comp_spec_Repeat : - forall (A B : Set)(eqda : EqDec A)(eqdb : EqDec B)(c1 : Comp A)(c2 : Comp B) P P1 P2, - comp_spec P c1 c2 -> - (forall a b, P a b -> (P1 a = P2 b)) -> - comp_spec P (Repeat c1 P1) (Repeat c2 P2). - - intuition. - destruct H. - intuition. - - exists (Repeat x (fun p => (P1 (fst p)) && (P2 (snd p)))). - intuition. - - simpl. - rewrite H1. - unfold marginal_l. - simpl. - - rewrite <- sumList_factor_constant_l. - - rewrite (sumList_filter_partition (fun p : A0 * B => P1 (fst p) && P2 (snd p))). - - assert (sumList - (filter (fun a : A0 * B => negb (P1 (fst a) && P2 (snd a))) - (getSupport x)) - (fun a : A0 * B => - indicator P1 r1 * - ratInverse (sumList (filter P1 (getSupport c1)) (evalDist c1)) * - (evalDist x a * - (if EqDec_dec bool_EqDec (eqb (fst a) r1) true then 1 else 0))) - == 0). - eapply sumList_0. - intuition. - eapply filter_In in H2. - intuition. - unfold indicator . - case_eq (P1 r1); intuition. - rewrite ratMult_1_l. - - destruct (EqDec_dec bool_EqDec (eqb (fst a) r1) true). - rewrite eqb_leibniz in e. - subst. - rewrite ratMult_1_r. - - eapply negb_true_iff in H5. - eapply andb_false_iff in H5. - intuition. - congruence. - specialize (H3 a). - intuition. - specialize (H0 _ _ H5). - congruence. - - repeat rewrite ratMult_0_r. - reflexivity. - - repeat rewrite ratMult_0_l. - reflexivity. - - rewrite H2. - clear H2. - rewrite <- ratAdd_0_r. - - eapply sumList_body_eq. - intuition. - eapply filter_In in H2. - intuition. - unfold indicator. - destruct (EqDec_dec bool_EqDec (eqb (fst a) r1) true). - rewrite eqb_leibniz in e. - subst. - eapply andb_true_iff in H5. - intuition. - rewrite H2. - simpl. - rewrite H6. - repeat rewrite ratMult_1_l. - repeat rewrite ratMult_1_r. - eapply ratMult_eqRat_compat; intuition. - eapply ratInverse_eqRat_compat. - eapply sumList_nz. - destruct a. - exists a. - intuition. - eapply filter_In. - intuition. - eapply getSupport_In_evalDist. - rewrite H1. - unfold marginal_l. - simpl. - eapply sumList_nz. - exists (a, b). - intuition. - simpl in H5. - rewrite eqb_refl in H5. - destruct ( EqDec_dec bool_EqDec true true). - rewrite ratMult_1_r in H5. - eapply getSupport_In_evalDist; - eauto. - intuition. - rewrite H1 in H5. - unfold marginal_l in H5. - simpl in H5. - eapply sumList_nz; eauto. - exists (a, b). - intuition. - simpl in H7. - rewrite eqb_refl in H7. - destruct (EqDec_dec bool_EqDec true true). - rewrite ratMult_1_r in H7. - eapply getSupport_In_evalDist; eauto. - intuition. - - - eapply eqRat_trans. - eapply sumList_body_eq. - intros. - eapply H1. - unfold marginal_l. - simpl. - rewrite sumList_summation. - - rewrite (sumList_filter_partition (fun p : A0 * B => P1 (fst p) && P2 (snd p))). - - assert (sumList - (filter (fun a0 : A0 * B => negb (P1 (fst a0) && P2 (snd a0))) - (getSupport x)) - (fun b : A0 * B => - sumList (filter P1 (getSupport c1)) - (fun a0 : A0 => - evalDist x b * - (if EqDec_dec bool_EqDec (eqb (fst b) a0) true then 1 else 0))) - == 0). - eapply sumList_0. - intuition. - eapply filter_In in H5. - intuition. - eapply sumList_0. - intuition. - eapply filter_In in H5. - intuition. - destruct (EqDec_dec bool_EqDec (eqb (fst a0) a1) true). - rewrite eqb_leibniz in e. - subst. - eapply negb_true_iff in H8. - eapply andb_false_iff in H8. - intuition. - congruence. - specialize (H3 _ H7). - specialize (H0 _ _ H3). - congruence. - eapply ratMult_0_r. - - rewrite H5. - clear H5. - rewrite <- ratAdd_0_r. - eapply sumList_body_eq. - intuition. - - eapply filter_In in H5. - intuition. - eapply andb_true_iff in H8. - intuition. - rewrite sumList_factor_constant_l. - symmetry. - rewrite <- ratMult_1_r at 1. - eapply ratMult_eqRat_compat. - reflexivity. - symmetry. - rewrite (@sumList_exactly_one _ (fst a0) ); intuition. - rewrite eqb_refl. - destruct (EqDec_dec bool_EqDec true true); intuition. - eapply filter_NoDup. - eapply getSupport_NoDup. - eapply filter_In. - intuition. - eapply getSupport_In_evalDist. - intuition. - rewrite H1 in H8. - unfold marginal_l in H8. - simpl in H8. - eapply sumList_nz; eauto. - econstructor. - intuition. - eauto. - eapply ratMult_0 in H10. - intuition. - eapply getSupport_In_evalDist. - eapply H7. - eauto. - rewrite eqb_refl in H11. - destruct ( EqDec_dec bool_EqDec true true); intuition. - case_eq ((eqb (fst a0) b)); intuition. - rewrite eqb_leibniz in H11. - intuition. - destruct (EqDec_dec bool_EqDec false true ). - discriminate. - intuition. - - repeat rewrite ratMult_0_r. - reflexivity. - - (* second marginal *) - - simpl. - rewrite H. - unfold marginal_r. - simpl. - - rewrite <- sumList_factor_constant_l. - - rewrite (sumList_filter_partition (fun p : A0 * B => P1 (fst p) && P2 (snd p))). - - assert (sumList - (filter (fun a : A0 * B => negb (P1 (fst a) && P2 (snd a))) - (getSupport x)) - (fun a : A0 * B => - indicator P2 r2 * - ratInverse (sumList (filter P2 (getSupport c2)) (evalDist c2)) * - (evalDist x a * - (if EqDec_dec bool_EqDec (eqb (snd a) r2) true then 1 else 0))) - == 0). - eapply sumList_0. - intuition. - eapply filter_In in H2. - intuition. - unfold indicator . - case_eq (P2 r2); intuition. - rewrite ratMult_1_l. - - destruct (EqDec_dec bool_EqDec (eqb (snd a) r2) true). - rewrite eqb_leibniz in e. - subst. - rewrite ratMult_1_r. - - eapply negb_true_iff in H5. - eapply andb_false_iff in H5. - intuition. - specialize (H3 a). - intuition. - specialize (H0 _ _ H5). - congruence. - congruence. - - repeat rewrite ratMult_0_r. - reflexivity. - - repeat rewrite ratMult_0_l. - reflexivity. - - rewrite H2. - clear H2. - rewrite <- ratAdd_0_r. - - eapply sumList_body_eq. - intuition. - eapply filter_In in H2. - intuition. - unfold indicator. - destruct (EqDec_dec bool_EqDec (eqb (snd a) r2) true). - rewrite eqb_leibniz in e. - subst. - eapply andb_true_iff in H5. - intuition. - rewrite H2. - simpl. - rewrite H6. - repeat rewrite ratMult_1_l. - repeat rewrite ratMult_1_r. - eapply ratMult_eqRat_compat; intuition. - eapply ratInverse_eqRat_compat. - eapply sumList_nz. - destruct a. - exists b. - intuition. - eapply filter_In. - intuition. - eapply getSupport_In_evalDist. - rewrite H. - unfold marginal_r. - simpl. - eapply sumList_nz. - exists (a, b). - intuition. - simpl in H5. - rewrite eqb_refl in H5. - destruct ( EqDec_dec bool_EqDec true true). - rewrite ratMult_1_r in H5. - eapply getSupport_In_evalDist; - eauto. - intuition. - rewrite H in H5. - unfold marginal_r in H5. - simpl in H5. - eapply sumList_nz; eauto. - exists (a, b). - intuition. - simpl in H7. - rewrite eqb_refl in H7. - destruct (EqDec_dec bool_EqDec true true). - rewrite ratMult_1_r in H7. - eapply getSupport_In_evalDist; eauto. - intuition. - - eapply eqRat_trans. - eapply sumList_body_eq. - intros. - eapply H. - unfold marginal_r. - simpl. - rewrite sumList_summation. - - rewrite (sumList_filter_partition (fun p : A0 * B => P1 (fst p) && P2 (snd p))). - - assert ( - sumList - (filter (fun a0 : A0 * B => negb (P1 (fst a0) && P2 (snd a0))) - (getSupport x)) - (fun b : A0 * B => - sumList (filter P2 (getSupport c2)) - (fun a0 : B => - evalDist x b * - (if EqDec_dec bool_EqDec (eqb (snd b) a0) true then 1 else 0))) - == 0). - eapply sumList_0. - intuition. - eapply filter_In in H5. - intuition. - eapply sumList_0. - intuition. - eapply filter_In in H5. - intuition. - destruct (EqDec_dec bool_EqDec (eqb (snd a0) a1) true). - rewrite eqb_leibniz in e. - subst. - eapply negb_true_iff in H8. - eapply andb_false_iff in H8. - intuition. - specialize (H3 _ H7). - specialize (H0 _ _ H3). - congruence. - congruence. - eapply ratMult_0_r. - - rewrite H5. - clear H5. - rewrite <- ratAdd_0_r. - eapply sumList_body_eq. - intuition. - - eapply filter_In in H5. - intuition. - eapply andb_true_iff in H8. - intuition. - rewrite sumList_factor_constant_l. - symmetry. - rewrite <- ratMult_1_r at 1. - eapply ratMult_eqRat_compat. - reflexivity. - symmetry. - rewrite (@sumList_exactly_one _ (snd a0) ); intuition. - rewrite eqb_refl. - destruct (EqDec_dec bool_EqDec true true); intuition. - eapply filter_NoDup. - eapply getSupport_NoDup. - eapply filter_In. - intuition. - eapply getSupport_In_evalDist. - intuition. - rewrite H in H8. - unfold marginal_r in H8. - simpl in H8. - eapply sumList_nz; eauto. - econstructor. - intuition. - eauto. - eapply ratMult_0 in H10. - intuition. - eapply getSupport_In_evalDist. - eapply H7. - eauto. - rewrite eqb_refl in H11. - destruct ( EqDec_dec bool_EqDec true true); intuition. - case_eq ((eqb (snd a0) b)); intuition. - rewrite eqb_leibniz in H11. - intuition. - destruct (EqDec_dec bool_EqDec false true ). - discriminate. - intuition. - - repeat rewrite ratMult_0_r. - reflexivity. - - (* predicate holds on all values in support of joint distribution *) - - unfold getSupport in H2. - fold getSupport in *. - eapply filter_In in H2. - intuition. - Qed. - - assert (option (TSet lambda * FreeList)). - apply None. - assert (Bvector lambda). - apply (oneVector lambda). - eapply comp_spec_Repeat. - comp_skip. - inline_first. - comp_skip. - comp_simp. - - assert( - comp_spec (fun x y => fst (fst y) = fst x /\ snd x = snd y) - (ret (a0, - (if a0 then true else false) && - negb - (eqb - (map - (fun x : Bvector lambda => - ESPADA_TSetRetrieve (getTSet a0) x) b0) - (map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) - (fst (t, l)) x) (snd (t, l)))))) - (ret (a0, b, - (if fst (a0, b) then true else false) && - negb - (eqb - (map - (fun x : Bvector lambda => - ESPADA_TSetRetrieve (getTSet (fst (a0, b))) x) b0) - (map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) l)))) - - ). - - eapply comp_spec_ret. - simpl; intuition. - eauto. - - intuition. - simpl in *; subst. - trivial. - - simpl in H2. - intuition. - subst. - simpl. - eapply comp_spec_eq_refl. - - rewrite H0. - clear H0. - - assert - ( - Pr - [z0 <-$ - Repeat - (z0 <-$ ESPADA_TSetSetup_trial t; - z1 <-$ - (tags <-$ - compMap (Bvector_EqDec lambda) - (fun x : Blist => ESPADA_TSetGetTag lambda F_bar (snd z0) x) l; - tSet <- getTSet (fst z0); - t_w <- - map (fun x : Bvector lambda => ESPADA_TSetRetrieve tSet x) tags; - t_w_correct <- - map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) l; - b <- negb (eqb t_w t_w_correct); - ret (if fst z0 then true else false) && b); - ret (z0, z1)) - (fun p : option (TSet lambda * FreeList) * Bvector lambda * bool => - isSome (fst (fst p))); ret snd z0 ] - ==Pr - [z0 <-$ ( - z0 <-$ - Repeat - (ESPADA_TSetSetup_trial t) - (fun p => - isSome (fst p)); - z1 <-$ - (tags <-$ - compMap (Bvector_EqDec lambda) - (fun x : Blist => ESPADA_TSetGetTag lambda F_bar (snd z0) x) l; - tSet <- getTSet (fst z0); - t_w <- - map (fun x : Bvector lambda => ESPADA_TSetRetrieve tSet x) tags; - t_w_correct <- - map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) l; - b <- negb (eqb t_w t_w_correct); - ret (if fst z0 then true else false) && b); - ret (z0, z1)); - ret snd z0 ] - ). - - comp_skip. - eapply (Repeat_factor_indep _ _ _ (fun x => isSome (fst x))). - - (* - Theorem ESPADA_TSetSetup_trial_wf : - forall t, - well_formed_comp (ESPADA_TSetSetup_trial t). - - intuition. - unfold ESPADA_TSetSetup_trial , ESPADA_TSet_Once_Games.ESPADA_TSetSetup_trial. - wftac. - eapply compFold_wf. - intuition. - unfold foldBody_option. - destruct a; wftac. - unfold ESPADA_TSetSetup_wLoop . - comp_simp. - eapply compFold_wf. - intuition. - unfold foldBody_option. - destruct a; wftac. - unfold ESPADA_TSetSetup_tLoop. - comp_simp. - - unfold maybeBindComp. - wftac. - eapply RndListElem.rndListElem_wf. - intuition. - destruct b4; wftac. - - Qed. - -*) - - eapply ESPADA_TSetSetup_trial_wf . - - assert (In (t, l, tt) (getSupport (x <-$ A; ret (x, tt)))). - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - specialize (@ESPADA_TSetSetup_Some_in_support lambda unit bigB bigS F F_bar - (x <-$ A; ret (x, tt)) failProb); intuition. - edestruct H2. - intuition. - - repeat simp_in_support. - unfold setupSuccess. - eapply leRat_trans. - Focus 2. - eapply failProb_correct. - eauto. - unfold ESPADA_TSetSetup_trial. - comp_skip. - - eapply eqRat_impl_leRat. - eapply evalDist_ret_eq. - unfold TSet. - unfold ESPADA_TSet_Once.TSet. - unfold TSet. - destruct ( @fst - (option - (prod (list (list (option (TSet_Record lambda)))) FreeList)) - (Bvector (posnatToNat lambda)) ); simpl; trivial. - eauto. - - destruct H1. - econstructor. - eapply filter_In. - eauto. - - intuition. - wftac. - eapply compMap_wf. - intuition. - unfold ESPADA_TSetGetTag. - wftac. - - rewrite H0. - clear H0. - - inline_first. - comp_skip. - reflexivity. - inline_first. - comp_simp. - comp_skip. - simpl. - reflexivity. - unfold getSupport in H0. - fold getSupport in H0. - eapply filter_In in H0. - intuition. - comp_simp. - eapply evalDist_ret_eq. - simpl. - simpl in H3. - destruct o; simpl. - reflexivity. - - simpl in H3. - discriminate. - Qed. - - Require Import RepeatCore. - - - Theorem ESPADA_TSet_Correct_Single_G1_equiv : - Pr[TrueSingle_G - (fun a : T lambda * list Blist => - TSetCorr_G0_RepeatBody a) A (fun x => snd x)] == - Pr[ESPADA_TSetCorr_G1 bigB bigS maxMatches F F_bar A ]. - - unfold TrueSingle_G, TSetCorr_G0_RepeatBody, ESPADA_TSetCorr_G1. - comp_skip. - comp_simp. - comp_inline_l. - unfold fst. - comp_skip. - reflexivity. - comp_simp. - - unfold snd. - - unfold ESPADA_TSetGetTag. - - inline_first. - - assert ( - Pr - [a <-$ (foreach (x in l)ret F_bar b x); - b0 <-$ - ret (o, - (if o then true else false) && - negb - (eqb - (map - (fun x : Bvector lambda => - ESPADA_TSetRetrieve (getTSet o) x) a) - (map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) l))); - ret (let (_, y) := b0 in y) ] - == - Pr - [a <-$ ret (map (F_bar b) l); - b0 <-$ - ret (o, - (if o then true else false) && - negb - (eqb - (map - (fun x : Bvector lambda => - ESPADA_TSetRetrieve (getTSet o) x) a) - (map - (fun x : list bool => - Array.arrayLookupList (list_EqDec bool_EqDec) t x) l))); - ret (let (_, y) := b0 in y) ] - ). - - inline_first. - comp_skip. - eapply comp_spec_eq_impl_eq. - eapply compMap_map. - - rewrite H1. - clear H1. - - comp_simp. - reflexivity. - Qed. - - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - forall (t : T lambda) (l : list Blist), - In (t, l) (getSupport A) -> - (length (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) <= - maxKeywords)%nat. - - Hypothesis maxMatches_correct : - forall (t : T lambda) (q : list Blist), - In (t, q) (getSupport A) -> maxMatch t maxMatches. - - Require Import PRF. - Require Import ESPADA_TSet_Once. - Local Open Scope rat_scope. - - Hypothesis A_wf : well_formed_comp A. - - Variable PRF_NA_Adv : Rat. - Hypothesis PRF_NA_Adv_correct_1 : - PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar - (list_EqDec bool_EqDec) (Bvector_EqDec lambda) - (ESPADA_TSet_PRF_A1 - (pair_EqDec - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (z <-$ A; [t, l]<-2 z; ret (t, l, (t, l)))) - (ESPADA_TSet_PRF_A2 bigB bigS F - (fun (s : T lambda * list Blist) - (p : option (TSet lambda) * list (Bvector lambda)) => - [t, q0]<-2 s; - [tSet_opt, tags]<-2 p; - tSet <- match tSet_opt with - | Some x => x - | None => nil - end; - ret (if tSet_opt then true else false) && - negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q0) - arrayLookupList (list_EqDec bool_EqDec) t x)))) - <= PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_2: - forall i, - PRF_NA_Advantage ({ 0 , 1 }^lambda) (RndF_R lambda bigB) F _ _ - (Hybrid.B1 nil _ _ (PRFI_A1 maxMatches A) i) - (Hybrid.B2 - (fun lsD => k <-$ {0, 1}^lambda; ret (map (F k) lsD)) - (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func (@RndF_range lambda bigB) _) nil lsD; ret lsR) _ - (PRFI_A2 lambda)) - <= PRF_NA_Adv. - - Theorem ESPADA_TSet_Correct : - AdvCor _ - (ESPADA_TSetSetup bigB bigS F F_bar) - (ESPADA_TSetGetTag lambda F_bar) - (ESPADA_TSetRetrieve) - A <= - (lambda / 1) * ((maxKeywords * (S maxMatches))^2 / 2 ^ lambda + - (maxKeywords / 1) * PRF_NA_Adv - + - PRF_NA_Adv) + (expRat failProb lambda). - - - rewrite ESPADA_TSet_Correct_G1_equiv. - unfold TSetCorr_G1. - rewrite (@TrueMult_impl_Repeat _ _ _ A A_wf TSetCorr_G0_RepeatBody (fun p => snd p) (fun p => isSome (fst p))); intuition. - - rewrite TrueSingle_impl_Mult. - eapply ratAdd_leRat_compat. - eapply ratMult_leRat_compat. - reflexivity. - - rewrite ESPADA_TSet_Correct_Single_G1_equiv. - - rewrite ESPADA_TSetCorr_once. - eapply ratAdd_leRat_compat. - reflexivity. - unfold ESPADA_TSetRetrieve. - reflexivity. - intuition. - trivial. - intuition. - intuition. - - eauto. - eauto. - reflexivity. - - unfold TSetCorr_G0_RepeatBody. - wftac. - eapply ESPADA_TSetSetup_trial_wf. - eapply compMap_wf. - intuition. - unfold ESPADA_TSetGetTag. - wftac. - - - destruct x. - assert (In (t, l, tt) (getSupport (z <-$ A; ret (z, tt)))). - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - specialize (@ESPADA_TSetSetup_Some_in_support lambda unit bigB bigS F F_bar - (z <-$ A; ret (z, tt)) failProb); intuition. - - edestruct H2. - intuition. - eapply leRat_trans. - Focus 2. - eapply failProb_correct. - repeat simp_in_support. - eauto. - comp_skip. - reflexivity. - unfold setupSuccess. - destruct x. - destruct o; - simpl; reflexivity. - eauto. - - destruct H1. - - econstructor. - eapply filter_In. - intuition. - unfold TSetCorr_G0_RepeatBody. - eapply getSupport_In_Seq. - eapply H1. - comp_simp. - eapply getSupport_In_Seq. - - unfold ESPADA_TSetGetTag. - simpl. - - eapply getSupport_In_evalDist. - intuition. - - assert (forall z, - evalDist - (compMap (Bvector_EqDec lambda) (fun x1 : Blist => ret F_bar x x1) l) z - == - evalDist (ret (map (F_bar x) l)) z). - - intuition. - eapply comp_spec_eq_impl_eq. - eapply compMap_map. - rewrite H4 in H3. - eapply getSupport_In_evalDist; [idtac | eauto]. - simpl. - left. - reflexivity. - simpl. - left. - reflexivity. - trivial. - - destruct a. - eapply leRat_trans. - Focus 2. - eapply failProb_correct. - eauto. - - unfold TSetCorr_G0_RepeatBody. - comp_inline_l. - comp_skip. - reflexivity. - comp_simp. - inline_first. - comp_irr_l. - eapply compMap_wf. - intuition. - unfold ESPADA_TSetGetTag. - wftac. - dist_compute. - Qed. - - Print Assumptions ESPADA_TSet_Correct. - -End ESPADA_TSet_Correct. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Correct_Once.v b/src/ESPADA/ESPADA_TSet_Correct_Once.v deleted file mode 100644 index c9bd8c8..0000000 --- a/src/ESPADA/ESPADA_TSet_Correct_Once.v +++ /dev/null @@ -1,5514 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import FCF. - -Require Import ESPADA_TSet_Once. -Require Import ESPADA_TSet_Once_Games. -Require Import TSet. -Require Import Array. -Require Import CompFold. -Require Import HasDups. -Require Import ESPADA_TSet_Correct_Once_Games. - -Local Open Scope list_scope. - - -Section ESPADA_TSet_Once_correct. - - Variable lambda : posnat. - Variable tSize bigB bigS maxMatches : posnat. - - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Hypothesis F_b_correct : - forall k b i v1 v2, - (b, v1, v2) = F k i -> - b < bigB. - - Variable A1 : Comp ((T lambda * list Blist)). - Hypothesis A1_wf : well_formed_comp A1. - - Definition maxMatch (t : T lambda) n := - forall p, - In p t -> - (length (snd p) <= n)%nat. - - Hypothesis maxMatches_correct : - forall t q, - In (t, q) (getSupport A1) -> - maxMatch t maxMatches. - - Definition ESPADA_TSetRetrieve := ESPADA_TSetRetrieve lambda maxMatches F. - Notation "'gInst' g" := (@g lambda bigB bigS maxMatches F F_bar A1) (at level 90). - Notation "'gInst1' g" := (@g lambda bigB bigS maxMatches F A1) (at level 90). - - Theorem G1_secure_1_equiv : - Pr[gInst ESPADA_TSetCorr_G1] == - Pr[ESPADA_TSet_1 bigB bigS F F_bar - ([t, l] <-$2 A1; ret (t, l, (t, l))) - (fun s p => [t, q] <-2 s; [tSet_opt, tags] <-2 p; - tSet <- match tSet_opt with | None => nil | Some x => x end; - ret ((if tSet_opt then true else false) && negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x))))]. - - intuition. - unfold ESPADA_TSetCorr_G1, ESPADA_TSet_1, ESPADA_TSetSetup_trial. - comp_simp. - do 3 (inline_first; comp_skip; comp_simp). - reflexivity. - eapply evalDist_ret_eq. - unfold getTSet_opt. - destruct x0; simpl. - unfold ESPADA_TSetRetrieve. - f_equal. - - trivial. - - Qed. - - Theorem G3_secure_9_equiv : - Pr[ESPADA_TSet_9 bigB bigS F - ([t, l] <-$2 A1; ret (t, l, (t, l))) - (fun s p => [t, q] <-2 s; [tSet_opt, tags] <-2 p; - tSet <- match tSet_opt with | None => nil | Some x => x end; - ret ((if tSet_opt then true else false) && negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x))))] - == Pr[ gInst1 ESPADA_TSetCorr_G3]. - - intuition. - unfold ESPADA_TSet_9, ESPADA_TSetCorr_G3. - comp_simp. - do 3 (inline_first; comp_skip; comp_simp). - reflexivity. - - unfold getTSet_opt. - destruct x0. - simpl. - reflexivity. - simpl. - intuition. - Qed. - - Theorem ESPADA_TSetCorr_G1_G3_close : - | Pr[gInst ESPADA_TSetCorr_G1] - Pr[gInst1 ESPADA_TSetCorr_G3] | - <= - PRF.PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar - _ _ - (ESPADA_TSet_PRF_A1 _ ([t, l] <-$2 A1; ret (t, l, (t, l)))) - (ESPADA_TSet_PRF_A2 bigB bigS F - (fun (s : (T lambda * list Blist)) - (p : option (TSet lambda) * list (Bvector lambda)) => - [t, q0] <-2 s; - [tSet_opt, tags]<-2 p; - tSet <- match tSet_opt with - | Some x => x - | None => nil - end; - ret ((if tSet_opt then true else false) && negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q0) - arrayLookupList (list_EqDec bool_EqDec) t x))))). - - intuition. - rewrite G1_secure_1_equiv. - rewrite <- G3_secure_9_equiv. - rewrite ESPADA_TSet_2_equiv. - rewrite ESPADA_TSet_3_1_equiv. - rewrite ESPADA_TSet_3_2_equiv. - rewrite ESPADA_TSet_3_equiv. - rewrite ESPADA_TSet_4_equiv. - rewrite ESPADA_TSet_6_equiv. - rewrite ESPADA_TSet_7_equiv. - rewrite ESPADA_TSet_8_1_equiv. - - rewrite <- ESPADA_TSet_9_equiv. - rewrite <- ESPADA_TSet_9_1_equiv. - rewrite <- ESPADA_TSet_8_equiv. - - eapply ESPADA_TSet_8_2_equiv. - apply F_bar. - - apply (fun x => (pos 1, pos 1)). - apply (fun x => (pos 1, pos 1)). - - intuition. - - apply (fun x => (pos 1, pos 1)). - apply (fun x => (pos 1, pos 1)). - apply (fun x => (pos 1, pos 1)). - Qed. - - Theorem ESPADA_TSetCorr_G3_G4_equiv : - Pr[gInst1 ESPADA_TSetCorr_G3] == Pr[gInst1 ESPADA_TSetCorr_G4]. - - intuition. - unfold ESPADA_TSetCorr_G3, ESPADA_TSetCorr_G4. - comp_simp. - comp_skip; comp_simp. - comp_skip. - comp_skip. - eapply evalDist_ret_eq. - f_equal. - - repeat rewrite map_map. - unfold combineKeywords. - - assert ( - (@map Blist (list (Vector.t bool (posnatToNat lambda))) - (fun x1 : Blist => - ESPADA_TSetRetrieve (@getTSet lambda x0) - (@nth (Bvector (posnatToNat lambda)) - (@lookupIndex Blist - (@EqDec_dec Blist (@list_EqDec bool bool_EqDec)) - (@app Blist - (@removeDups Blist (@list_EqDec bool bool_EqDec) l) - (@removePresent Blist - (@EqDec_dec Blist (@list_EqDec bool bool_EqDec)) - (@removeDups Blist (@list_EqDec bool bool_EqDec) l) - (@toW lambda t))) x1 O) x - (@Vector.const bool false (posnatToNat lambda)))) l) - - = - (foreach (x1 in l) - ESPADA_TSetRetrieve (getTSet x0) - (nth - (lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x1 0) x (Vector.const false lambda))) - ). - - eapply map_ext_in. - intuition. - - rewrite lookupIndex_app. - trivial. - apply (fun x => (pos 1, pos 1)). - eapply in_removeDups. - trivial. - - unfold ESPADA_TSetRetrieve in *. - rewrite H2. - clear H2. - - assert ( - (@map (Bvector (posnatToNat lambda)) - (list (Vector.t bool (posnatToNat lambda))) - (fun x1 : Bvector (posnatToNat lambda) => - ESPADA_TSetRetrieve (@getTSet lambda x0) x1) - (@firstn (Bvector (posnatToNat lambda)) - (@length Blist - (@removeDups Blist (@list_EqDec bool bool_EqDec) l)) x)) - = - (foreach (x1 in (removeDups _ l)) - ESPADA_TSetRetrieve (getTSet x0) - (nth - (lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x1 0) x - (Vector.const false lambda))) - ). - - eapply map_eq. - - eapply list_pred_impl. - - Theorem lp_G3_G4_equiv : - forall q x, - NoDup q -> - length x >= length q -> - list_pred - (fun (a : Bvector lambda) (b : list bool) => - a = - nth - (lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - q b 0) x - (Vector.const false lambda)) - (firstn (length q) x) - q. - - induction q; destruct x; intuition; simpl in *. - econstructor. - econstructor. - omega. - - econstructor. - destruct (EqDec_dec (list_EqDec bool_EqDec) a a); trivial. - congruence. - - inversion H; clear H; subst. - - eapply list_pred_impl'. - eapply IHq. - trivial. - omega. - intuition. - subst. - destruct ( EqDec_dec (list_EqDec bool_EqDec) b a). - subst. - intuition. - - trivial. - - Qed. - - eapply lp_G3_G4_equiv. - eapply removeDups_NoDup. - eapply compMap_length in H0. - eapply Nat.le_stepr. - Focus 2. - symmetry. - eapply H0. - unfold combineKeywords. - rewrite app_length. - rewrite <- plus_0_r at 1. - eapply plus_le_compat; intuition. - - intuition. - subst. - trivial. - - unfold ESPADA_TSetRetrieve in *. - rewrite H2. - clear H2. - - case_eq (eqb (foreach (x1 in l) - ESPADA_TSet_Correct_Once_Games.ESPADA_TSetRetrieve lambda maxMatches F (getTSet x0) - (nth - (lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x1 0) x - (Vector.const false lambda))) - (foreach (x1 in l)arrayLookupList (list_EqDec bool_EqDec) t x1)); intuition. - rewrite eqb_leibniz in H2. - - unfold ESPADA_TSetRetrieve in *. - eapply map_eq_subset in H2. - rewrite H2. - repeat rewrite eqb_refl. - trivial. - intuition. - eapply removeDups_in. - trivial. - - case_eq (eqb - (@map (list bool) (list (Vector.t bool (posnatToNat lambda))) - (fun x1 : list bool => - ESPADA_TSet_Correct_Once_Games.ESPADA_TSetRetrieve lambda maxMatches F - (@getTSet lambda x0) - (@nth (Bvector (posnatToNat lambda)) - (@lookupIndex (list bool) - (@EqDec_dec (list bool) (@list_EqDec bool bool_EqDec)) - (@removeDups (list bool) (@list_EqDec bool bool_EqDec) l) - x1 O) x (@Vector.const bool false (posnatToNat lambda)))) - (@removeDups Blist (@list_EqDec bool bool_EqDec) l)) - (@map (list bool) (list (Bvector (posnatToNat lambda))) - (fun x1 : list bool => - @arrayLookupList (list bool) (Bvector (posnatToNat lambda)) - (@list_EqDec bool bool_EqDec) t x1) - (@removeDups Blist (@list_EqDec bool bool_EqDec) l)) - ); intuition. - - rewrite eqb_leibniz in H3. - eapply map_eq_subset in H3. - rewrite H3 in H2. - rewrite eqb_refl in H2. - discriminate. - intuition. - eapply in_removeDups; eauto. - Qed. - - - Definition tSet_6_Corr_5_equiv (t1 : TSet lambda)(t2 : TSetCorr_5 lambda) := - list_pred - (fun x y => - list_pred - (fun a b => - match a with - | None => b = None - | Some z1 => - match b with - | None => False - | Some z2 => - fst z1 = fst z2 /\ - snd z1 = snd (snd z2) - end - end) x y) - t1 t2. - - Theorem ESPADA_TSetSetup_6_Corr_5_tLoop_equiv : - forall t1 t2, - tSet_6_Corr_5_equiv t1 t2 -> - forall a1 a2 f, - (let '(q0, tag, i, len, s_i, b, bigL, bigK) := a2 in - (i, len, s_i, b, bigL, bigK) = a1) -> - comp_spec (fun a b => - match a with - | None => b = None - | Some x => - match b with - | None => False - | Some y => - snd x = snd y /\ - tSet_6_Corr_5_equiv (fst x) (fst y) - end - end) - (ESPADA_TSetSetup_tLoop_6 (t1, f) a1) - (ESPADA_TSetSetup_tLoop_Corr_5 (t2, f) a2). - - intuition. - unfold ESPADA_TSetSetup_tLoop_6, ESPADA_TSet_Once_Games.ESPADA_TSetSetup_tLoop_6 , ESPADA_TSetSetup_tLoop_Corr_5. - pairInv. - unfold setLet. - comp_skip. - apply None. - apply None. - destruct b4. - eapply comp_spec_ret; intuition. - unfold fst. - unfold tSetUpdate, tSetUpdate_Corr_5. - unfold tSet_6_Corr_5_equiv. - - eapply list_pred_listReplace. - eapply H. - - eapply list_pred_listReplace. - - eapply list_pred_nth. - eapply H. - econstructor. - - - unfold fst, snd. - intuition. - trivial. - - econstructor. - - eapply comp_spec_ret; intuition. - - Qed. - - Theorem ESPADA_TSetSetup_6_Corr_5_equiv : - forall ls1 ls2, - list_pred (fun x y => - let '(q0, tag, i, len, s_i, b, bigL, bigK) := y in - (i, len, s_i, b, bigL, bigK) = x) - ls1 ls2 -> - forall t1 t2 f, - tSet_6_Corr_5_equiv t1 t2 -> - comp_spec (fun a b => - match a with - | None => b = None - | Some x => - match b with - | None => False - | Some y => - snd x = snd y /\ - tSet_6_Corr_5_equiv (fst x) (fst y) - end - end) - (compFold _ - (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (t1, f)) ls1) - (compFold _ - (foldBody_option _ - (@ESPADA_TSetSetup_tLoop_Corr_5 lambda)) (Some (t2, f)) ls2). - - induction 1; intuition. simpl in *. - - eapply comp_spec_ret; intuition. - - eapply comp_spec_eq_trans_l. - eapply compFold_cons. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_eq_symm. - eapply compFold_cons. - comp_skip. - eapply None. - eapply None. - unfold foldBody_option. - eapply ESPADA_TSetSetup_6_Corr_5_tLoop_equiv. - - trivial. - intuition. - - simpl in H4. - destruct a. - destruct b. - destruct p, p0. - intuition. - simpl in H5. - subst. - eapply IHlist_pred. - trivial. - intuition. - subst. - - - eapply comp_spec_consequence. - eapply (@compFold_spec' _ _ _ _ (fun _ _ c d => c = None /\ d = None)). - eapply list_pred_length_eq. - eapply H0. - intuition. - - intuition. - subst. - unfold foldBody_option. - eapply comp_spec_ret; intuition. - - intuition. - intuition. - subst. - trivial. - - Qed. - - Theorem ESPADA_TSetCorr_G4_G5_equiv : - Pr[ gInst1 ESPADA_TSetCorr_G4] == Pr[ESPADA_TSetCorr_G5 bigB bigS maxMatches F A1]. - - intuition. - unfold ESPADA_TSetCorr_G4, ESPADA_TSetCorr_G5. - comp_simp. - comp_skip; comp_simp. - comp_skip. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply ESPADA_TSetSetup_6_Corr_5_equiv. - - assert (length (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) = - length x). - eapply compMap_length in H0. - intuition. - - eapply list_pred_flatten_both. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_zip_r. - eapply list_pred_zip_l. - eapply list_pred_I; intuition. - - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_map_r. - eapply list_pred_I; intuition. - - eapply list_pred_zip_l. - eapply list_pred_I; intuition. - - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_I; intuition. - - eapply list_pred_I; intuition. - - eapply list_pred_map_l. - eapply list_pred_eq. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_I; intuition. - - eapply list_pred_eq. - eapply list_pred_map_l. - eapply list_pred_I; intuition. - - eapply list_pred_map_l. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_I; intuition. - - eapply list_pred_I; intuition. - - eapply list_pred_map_l. - eapply list_pred_eq. - intuition. - - Ltac exdest := - match goal with - | [H : exists _ : _, _ |- _] => destruct H - end. - - repeat exdest; intuition; simpl in *; subst. - repeat exdest; intuition; simpl in *; subst. - destruct b0. - simpl in *. subst. - - unfold numberedMap. - destruct p. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - rewrite H12. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst. - simpl. - - remember ( F b0 a0) as z. - destruct z. - destruct p. - trivial. - - econstructor. - - eapply comp_spec_ret; intuition. - destruct a. - destruct b. - destruct p, p0. - simpl in *. - intuition; subst. - f_equal. - f_equal. - - eapply map_ext_in. - intuition. - - Theorem ESPADA_TSetRetrieve_Corr_5_arrayLookup_eq: - forall t1 t2, - tSet_6_Corr_5_equiv t1 t2 -> - forall n a, - arrayLookupOpt (Bvector_EqDec lambda) (nth n t1 nil) a = - match (arrayLookupOpt (Bvector_EqDec lambda) (nth n t2 nil) a) with - | Some (x, y) => Some y - | None => None - end. - - intuition. - - eapply (@list_pred_nth _ _ _ (fun x y => arrayLookupOpt _ x a = - match arrayLookupOpt (Bvector_EqDec lambda) y a with - | Some (_, y) => Some y - | None => None - end)). - eapply list_pred_impl. - eapply H. - intuition. - - specialize (@list_pred_arrayLookupOpt _ _ _ (fun z1 z2 => z1 = snd (z2)) _ _ _ H1 a); intuition. - case_eq (arrayLookupOpt (Bvector_EqDec lambda) a0 a ); intuition; - rewrite H3 in H2. - destruct H2. - intuition. - subst. - rewrite H4. - destruct x. - trivial. - - rewrite H2. - trivial. - - simpl. - trivial. - - Qed. - - - Theorem ESPADA_TSetRetrieve_tLoop_Corr_5_equiv : - forall i t1 t2 a, - tSet_6_Corr_5_equiv t1 t2 -> - ESPADA_TSetRetrieve_tLoop lambda F t1 a i = ESPADA_TSetRetrieve_tLoop_Corr_5 F t2 a i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop, ESPADA_TSetRetrieve_tLoop_Corr_5. - remember (F a i). - comp_simp. - rewrite (@ESPADA_TSetRetrieve_Corr_5_arrayLookup_eq _ t2). - case_eq ( arrayLookupOpt (Bvector_EqDec lambda) (nth n t2 nil) b0); intuition. - destruct p. - trivial. - trivial. - Qed. - - Theorem ESPADA_TSetRetrieve_h_Corr_5_equiv : - forall fuel i t1 t2 a, - tSet_6_Corr_5_equiv t1 t2 -> - ESPADA_TSetRetrieve_h lambda F t1 a i fuel = ESPADA_TSetRetrieve_h_Corr_5 F t2 a i fuel. - - induction fuel; intuition; simpl in *. - rewrite (@ESPADA_TSetRetrieve_tLoop_Corr_5_equiv _ _ t2). - remember (ESPADA_TSetRetrieve_tLoop_Corr_5 F t2 a i). - destruct o; trivial. - comp_simp. - f_equal. - destruct b; trivial. - eauto. - trivial. - Qed. - - Theorem ESPADA_TSetRetrieve_Corr_5_equiv : - forall t1 t2 a, - tSet_6_Corr_5_equiv t1 t2 -> - ESPADA_TSetRetrieve t1 a = ESPADA_TSetRetrieve_Corr_5 maxMatches F t2 a. - - intuition. - unfold ESPADA_TSetRetrieve, ESPADA_TSetRetrieve_Corr_5, ESPADA_TSet_Correct_Once_Games.ESPADA_TSetRetrieve. - unfold ESPADA_TSet.ESPADA_TSetRetrieve. - eapply ESPADA_TSetRetrieve_h_Corr_5_equiv. - trivial. - - Qed. - - eapply ESPADA_TSetRetrieve_Corr_5_equiv. - trivial. - intuition. - subst. - trivial. - Qed. - - - Theorem ESPADA_TSetCorr_G5_le_G6 : - Pr[ gInst1 ESPADA_TSetCorr_G5 ] <= Pr[ gInst1 ESPADA_TSetCorr_G6 ]. - - intuition. - unfold ESPADA_TSetCorr_G5, ESPADA_TSetCorr_G6. - comp_simp. - comp_skip; comp_simp. - comp_skip. - comp_skip. - eapply comp_spec_impl_le. - eapply comp_spec_ret. - intuition. - destruct x0; simpl in *; trivial. - - destruct ( labelCollision_6 (fst p)); trivial. - - Qed. - - - Definition ESPADA_TSetSetup_tLoop_Corr_6_1 (tSet : TSetCorr_7 lambda) - (e : Blist * Bvector lambda * nat * nat * Vector.t bool lambda * nat * Bvector lambda * - Vector.t bool (S lambda)) := - (let '(q, tag, i, len, s_i, b, bigL, bigK) := e in - bet <- (if eq_nat_dec (S i) len then false else true); - newRecord <- (bigL, ((q, tag, i, len, s_i, b, bigK), (Vector.cons bool bet lambda s_i xor bigK))); - newBucket <- (nth b tSet nil) ++ (newRecord :: nil); - ret listReplace tSet b newBucket nil - ). - - Definition ESPADA_TSetCorr_G6_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <-$ compFold _ ( ESPADA_TSetSetup_tLoop_Corr_6_1) nil (flatten ts_recsList); - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_7 maxMatches F tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret (labelCollision_7 ( tSet) || bad). - - - Require Import Permutation. - Require Import RndListElem. - - Theorem TSetSetup_5_6_1_spec : - forall d p x, - (forall n, Permutation (getSomes (nth n (fst p) nil)) (nth n d nil)) /\ - (forall b n, In n (nth b (snd p) nil) -> nth n (nth b (fst p) nil) None = None) /\ - (forall n, NoDup (nth n (snd p) nil)) -> - - comp_spec - (fun (a : option (list (list (option (TSetCorr_5_Record lambda))) * FreeList)) - (b6 : list (list (TSetCorr_5_Record lambda))) => - match a with - | Some p0 => - (forall n, Permutation (getSomes (nth n (fst p0) nil)) (nth n b6 nil)) /\ - - (forall b n, In n (nth b (snd p0) nil) -> nth n (nth b (fst p0) nil) None = None) /\ - (forall n, NoDup (nth n (snd p0) nil)) - | None => True - end) (ESPADA_TSetSetup_tLoop_Corr_5 p x) - (ESPADA_TSetSetup_tLoop_Corr_6_1 d x). - - intros. - unfold ESPADA_TSetSetup_tLoop_Corr_5, ESPADA_TSetSetup_tLoop_Corr_6_1. - comp_simp. - comp_irr_l. - eapply RndListElem.rndListElem_wf. - - destruct a. - eapply comp_spec_ret; intros. - unfold tSetUpdate_Corr_5. - - unfold fst. - intuition. - - destruct (eq_nat_dec n3 n). subst. - repeat rewrite nth_listReplace_eq. - - eapply perm_trans. - eapply listReplace_getSomes_Permutation. - eapply H. - eapply rndListElem_support. - trivial. - trivial. - eapply Permutation_cons_append. - - rewrite nth_listReplace_ne; intuition. - rewrite nth_listReplace_ne; intuition. - - unfold snd in *. - destruct (eq_nat_dec b2 n); subst. - rewrite nth_listReplace_eq in *. - destruct (eq_nat_dec n3 n2); subst. - - exfalso. - eapply removeFirst_NoDup_not_in. - eapply H2. - eauto. - - rewrite nth_listReplace_ne; intuition. - eapply H. - eapply removeFirst_in_iff. - eapply H4. - - rewrite nth_listReplace_ne in H4; intuition. - rewrite nth_listReplace_ne; intuition. - unfold snd. - destruct (eq_nat_dec n3 n); subst. - rewrite nth_listReplace_eq. - eapply removeFirst_NoDup. - eauto. - rewrite nth_listReplace_ne. - eauto. - intuition. - - eapply comp_spec_ret; intuition. - Qed. - - Theorem perm_labelIn_7_eq : - forall ls1 ls2, - Permutation ls1 ls2 -> - forall b, - (@labelIn_7 lambda) ls1 b = labelIn_7 ls2 b. - - induction 1; intuition; simpl in *. - - destruct x. - case_eq (eqbBvector b b0); intuition. - - destruct y. - destruct x. - case_eq (eqbBvector b b0); intuition. - destruct (eqbBvector b b1); intuition. - - rewrite IHPermutation1. - eapply IHPermutation2. - - Qed. - - Theorem perm_bucketCollision_7_eq : - forall ls1 ls2, - Permutation ls1 ls2 -> - (@bucketLabelCollision_7 lambda) ls1 = - bucketLabelCollision_7 ls2. - - induction 1; intuition; simpl in *. - - destruct x. - erewrite perm_labelIn_7_eq. - f_equal. - intuition. - trivial. - - destruct y. - destruct x. - case_eq (eqbBvector b b0); intuition. - simpl. - apply eqbBvector_sound in H. - subst. - rewrite eqbBvector_complete. - simpl. - trivial. - - case_eq (eqbBvector b0 b); intuition. - apply eqbBvector_sound in H0. - subst. - rewrite eqbBvector_complete in H. - discriminate. - - destruct (labelIn_7 l b); destruct (labelIn_7 l b0); simpl; trivial. - - rewrite IHPermutation1. - eapply IHPermutation2. - Qed. - - Theorem noBucketCollision_arrayLookup_eq : - forall ls1 ls2, - Permutation ls1 ls2 -> - (@bucketLabelCollision_7 lambda) ls1 = false -> - forall b, - arrayLookup _ ls1 b = arrayLookup _ ls2 b. - - induction 1; intuition; simpl in *. - - destruct x. - case_eq (eqbBvector b b0); intuition. - eapply IHPermutation. - eapply orb_false_iff in H0. - intuition. - - destruct y. - destruct x. - case_eq (eqbBvector b0 b1); intuition; - rewrite H1 in H. - simpl in *. discriminate. - apply orb_false_iff in H. - intuition. - apply orb_false_iff in H3. - intuition. - case_eq (eqbBvector b b0); intuition. - apply eqbBvector_sound in H3. - subst. - case_eq (eqbBvector b0 b1); intuition. - apply eqbBvector_sound in H3. - subst. - rewrite eqbBvector_complete in H1. - discriminate. - - rewrite H2. - eapply IHPermutation2. - erewrite perm_bucketCollision_7_eq; eauto. - eapply Permutation_sym. - trivial. - Qed. - - Theorem TSetRetrieve_tLoop_Corr_5_7_arrayLookup_eq : - forall ls1 ls2 b0, - Permutation (getSomes ls1) ls2 -> - (@bucketLabelCollision_7 lambda) ls2 = false -> - arrayLookupOpt _ ls1 b0 = - arrayLookup _ ls2 b0. - - intuition. - rewrite arrayLookupOpt_getSomes_eq. - eapply noBucketCollision_arrayLookup_eq ; intuition. - erewrite perm_bucketCollision_7_eq; eauto. - Qed. - - Theorem TSetRetrieve_tLoop_Corr_5_7_eq : - forall i t1 t2 a, - labelCollision_7 t2 = false -> - (forall n, Permutation (getSomes (nth n t1 nil)) (nth n t2 nil)) -> - ESPADA_TSetRetrieve_tLoop_Corr_5 F t1 a i = ESPADA_TSetRetrieve_tLoop_Corr_7 F t2 a i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_5, ESPADA_TSetRetrieve_tLoop_Corr_7. - - remember (F a i). - comp_simp. - rewrite (@TSetRetrieve_tLoop_Corr_5_7_arrayLookup_eq _ (nth n t2 nil)). - trivial. - - eauto. - - Theorem labelCollision_7_false_all : - forall ls, - (@labelCollision_7 lambda) ls = false -> - forall n, - bucketLabelCollision_7 (nth n ls nil) = false. - - induction ls; intuition; simpl in *. - destruct n; trivial. - - apply orb_false_iff in H. - intuition. - destruct n; intuition. - - Qed. - - eapply labelCollision_7_false_all. - trivial. - Qed. - - Theorem labelCollision_6_false_all_if : - forall ls, - (forall n, - (@bucketLabelCollision_6 lambda) (nth n ls nil) = false) -> - labelCollision_6 ls = false. - - induction ls; intuition; simpl in *. - - eapply orb_false_iff. - intuition. - specialize (H O). - simpl in *. - intuition. - - eapply IHls. - intuition. - specialize (H (S n)). - simpl in *. - intuition. - - Qed. - - Theorem TSetRetrieve_Corr_5_7_eq : - forall fuel i t1 t2 a, - labelCollision_7 t2 = false -> - (forall n, Permutation (getSomes (nth n t1 nil)) (nth n t2 nil)) -> - ESPADA_TSetRetrieve_h_Corr_5 F t1 a i fuel = ESPADA_TSetRetrieve_h_Corr_7 F t2 a i fuel. - - induction fuel; intuition; simpl in *. - - assert (ESPADA_TSetRetrieve_tLoop_Corr_5 F t1 a i = - ESPADA_TSetRetrieve_tLoop_Corr_7 F t2 a i). - eapply TSetRetrieve_tLoop_Corr_5_7_eq; intuition. - rewrite H2. - clear H2. - - remember (ESPADA_TSetRetrieve_tLoop_Corr_7 F t2 a i). - destruct o. - comp_simp. - f_equal. - destruct b; trivial. - eapply IHfuel; intuition. - trivial. - Qed. - - Theorem ESPADA_TSetCorr_G6_1_equiv : - Pr[ gInst1 ESPADA_TSetCorr_G6] <= Pr[ ESPADA_TSetCorr_G6_1]. - - unfold ESPADA_TSetCorr_G6, ESPADA_TSetCorr_G6_1. - - intuition. - comp_simp. - comp_skip; comp_simp. - comp_skip. - - eapply comp_spec_impl_le. - comp_skip. - eapply (@compFold_spec _ _ _ - (fun ls a b => - match a with - | None => True - | Some p => - (forall n, Permutation (getSomes (nth n (fst p) nil)) (nth n b nil)) /\ - - (forall b n, In n (nth b (snd p) nil) -> nth n (nth b (fst p) nil) None = None) /\ - (forall n, NoDup (nth n (snd p) nil)) - end)). - - intuition. - simpl. - destruct n; - simpl; - econstructor. - - simpl. - destruct b; - simpl; - destruct n; trivial. - - simpl. - destruct (lt_dec n (length (initFree bigB bigS))). - eapply initFree_correct. - apply (fun x => (pos 1, pos 1)). - eapply nth_In. - eapply l0. - rewrite nth_overflow. - econstructor. - omega. - - intuition. - destruct c. - unfold foldBody_option. - eapply TSetSetup_5_6_1_spec. - - intuition. - unfold foldBody_option. - eapply comp_spec_ret; intuition. - - simpl in *. - eapply comp_spec_ret; intuition. - destruct a; simpl in *. - intuition. - - case_eq (labelCollision_7 b); intuition. - rewrite orb_false_l. - - eapply negb_true_iff. - - assert (labelCollision_6 (fst p) = false). - - eapply labelCollision_6_false_all_if. - intuition. - eapply labelCollision_7_false_all in H6. - - Theorem labelIn_6_7_eq : - forall ls a, - (@labelIn_6 lambda) ls a = labelIn_7 (getSomes ls) a. - - induction ls; intuition; simpl in *. - destruct a. - destruct t. - simpl. - destruct (eqbBvector a0 b); intuition. - eauto. - - Qed. - - Theorem bucketLabelCollision_6_7_eq : - forall ls, - (@bucketLabelCollision_6 lambda) ls = bucketLabelCollision_7 (getSomes ls). - - induction ls; intuition; simpl in *. - destruct a. - destruct t. - simpl. - f_equal. - eapply labelIn_6_7_eq. - - intuition. - intuition. - Qed. - - rewrite bucketLabelCollision_6_7_eq. - erewrite perm_bucketCollision_7_eq . - eauto. - eauto. - - rewrite H8 in H4. - rewrite orb_false_l in H4. - eapply negb_true_iff in H4. - - eapply eqb_false_iff . - eapply eqb_false_iff in H4. - intuition. - eapply H4. - - rewrite <- H9. - eapply map_ext_in. - intuition. - - eapply TSetRetrieve_Corr_5_7_eq; intuition. - - discriminate. - - Qed. - - - Theorem ESPADA_TSetCorr_G6_1_G7_equiv : - Pr[ ESPADA_TSetCorr_G6_1] == Pr[ ESPADA_TSetCorr_G7 maxMatches F A1 ]. - - intuition. - unfold ESPADA_TSetCorr_G6_1, ESPADA_TSetCorr_G7. - comp_simp. - comp_skip; comp_simp. - comp_skip. - comp_irr_l. - eapply compFold_wf. - intuition. - unfold ESPADA_TSetSetup_tLoop_Corr_6_1. - wftac. - eapply evalDist_ret_eq. - - eapply (@compFold_support _ _ _ (fun t ls1 ls2 => t = fold_left (@ESPADA_TSetSetup_tLoop_Corr_7 lambda) ls1 nil)) in H1. - subst. - trivial. - trivial. - - intuition. - subst. - unfold ESPADA_TSetSetup_tLoop_Corr_6_1 in *. - unfold setLet in *. - repeat simp_in_support. - rewrite fold_left_app. - - trivial. - - Qed. - - Theorem ESPADA_TSetCorr_G6_G7_equiv : - Pr[ gInst1 ESPADA_TSetCorr_G6 ] <= Pr[ ESPADA_TSetCorr_G7 maxMatches F A1 ]. - - intuition. - rewrite ESPADA_TSetCorr_G6_1_equiv. - rewrite ESPADA_TSetCorr_G6_1_G7_equiv . - intuition. - - Qed. - - Definition labelCollision_7_1 (ls : list (list (TSetCorr_8_Record lambda))) := - labelCollision_8 lambda (flatten ls). - - Definition TSetCorr_7_1 := list (list (TSetCorr_8_Record lambda)). - - Definition ESPADA_TSetSetup_tLoop_Corr_7_1 (tSet : TSetCorr_7_1) - (e : Blist * Bvector lambda * nat * nat * Vector.t bool lambda * nat * Bvector lambda * - Vector.t bool (S lambda)) := - (let '(q, tag, i, len, s_i, b, bigL, bigK) := e in - bet <- (if eq_nat_dec (S i) len then false else true); - newRecord <- ((b, bigL), ((q, tag, i, len, s_i, b, bigK), (Vector.cons bool bet lambda s_i xor bigK))); - newBucket <- (nth b tSet nil) ++ ( newRecord :: nil); - listReplace tSet b newBucket nil - ). - - - Definition ESPADA_TSetRetrieve_tLoop_Corr_7_1 (tSet : TSetCorr_7_1) stag i := - [b, L, K] <-3 F stag i; - x <- nth b tSet nil; - t <- arrayLookup _ x (b, L); - match t with - | None => None - | Some p => - [_, u] <-2 p; - v <- u xor K; - bet <- Vector.hd v; - s <- Vector.tl v; - Some (s, bet) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_7_1 (tSet : TSetCorr_7_1) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_7_1 tSet stag i) with - | Some (v, bet) => - v :: (if (bet) then (ESPADA_TSetRetrieve_h_Corr_7_1 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_7_1 tSet stag := - ESPADA_TSetRetrieve_h_Corr_7_1 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G7_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- fold_left ( ESPADA_TSetSetup_tLoop_Corr_7_1) (flatten ts_recsList) nil; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_7_1 tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret (labelCollision_7_1 ( tSet) || bad). - - - Theorem ESPADA_TSetCorr_G7_1_equiv : - Pr[ESPADA_TSetCorr_G7 maxMatches F A1] == Pr[ESPADA_TSetCorr_G7_1]. - - intuition. - unfold ESPADA_TSetCorr_G7, ESPADA_TSetCorr_G7_1. - comp_simp. - comp_skip; comp_simp. - comp_skip. - - Theorem labelCollision_7_1_eq : - forall ls1 ls2, - list_pred - (fun a b => - list_pred - (fun x y => - fst x = snd (fst y) /\ snd x = snd y) a b) - ls1 ls2 -> - (forall n a, In a (nth n ls2 nil) -> fst (fst a) = n) -> - labelCollision_7 ls1 = - labelCollision_7_1 ls2. - - induction ls1 using rev_ind; intuition; simpl in *. - inversion H; clear H; subst. - - unfold labelCollision_7_1. - simpl. - trivial. - - specialize (@ls_last_exists _ ls2 (length ls1)); intuition. - edestruct H1. - eapply list_pred_length_eq in H. - rewrite app_length in H. - simpl in H. - rewrite plus_comm in H. - simpl in *. - intuition. - destruct H2. - intuition. - subst. - - Theorem labelCollision_7_app : - forall ls1 ls2, - (@labelCollision_7 lambda) (ls1 ++ ls2) = - labelCollision_7 ls1 || labelCollision_7 ls2. - - induction ls1; intuition; simpl in *. - rewrite IHls1. - rewrite orb_assoc. - trivial. - - Qed. - - - Theorem labelIn_8_app_l_false : - forall Z ls1 ls2 y, - @labelIn_8 lambda Z ls2 y = false -> - labelIn_8 lambda (ls1 ++ ls2) y = labelIn_8 lambda ls1 y. - - induction ls1; intuition; simpl in *. - destruct ( eqbPair nat_EqDec (Bvector_EqDec lambda) y (a, b0)); intuition. - - Qed. - - Theorem labelIn_8_app_r_true : - forall Z ls1 ls2 y, - @labelIn_8 lambda Z ls2 y = true -> - labelIn_8 lambda (ls1 ++ ls2) y = true. - - induction ls1; intuition; simpl in *. - destruct ( eqbPair nat_EqDec (Bvector_EqDec lambda) y (a, b0)); intuition. - - Qed. - - Theorem labelCollision_8_app : - forall Z ls1 ls2, - (forall x, labelIn_8 lambda ls1 x = true -> @labelIn_8 lambda Z ls2 x = false) -> - labelCollision_8 lambda (ls1 ++ ls2) = - labelCollision_8 lambda ls1 || labelCollision_8 lambda ls2. - - induction ls1; intuition; simpl in *. - - rewrite IHls1. - rewrite orb_assoc. - f_equal. - f_equal. - specialize (H0 (a, b0)). - eapply labelIn_8_app_l_false. - eapply H0. - unfold eqbPair. - repeat rewrite eqb_refl. - trivial. - - intuition. - eapply H0. - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) x (a, b0)); intuition. - - Qed. - - Theorem labelCollision_7_1_skipn : - forall ls n, - (forall j j' b, - j' > j -> - labelIn_8 lambda (nth j ls nil) b = true -> - labelIn_8 lambda (flatten (skipn j' ls)) b = false) -> - labelCollision_7_1 ls = - labelCollision_7_1 (firstn n ls) || labelCollision_7_1 (skipn n ls). - - induction ls; intuition; simpl in *. - - unfold labelCollision_7_1. - simpl. - rewrite firstn_nil. - - - rewrite skipn_nil. - simpl. - trivial. - - unfold labelCollision_7_1 in *. - simpl. - destruct n; simpl in *. - trivial. - rewrite labelCollision_8_app. - rewrite labelCollision_8_app. - rewrite <- orb_assoc. - f_equal. - eapply IHls. - intuition. - specialize (H (S j) (S j')%nat b). - simpl in *. - intuition. - - intuition. - - Theorem labelIn_8_app_false : - forall Z ls1 ls2 a, - @labelIn_8 lambda Z (ls1 ++ ls2) a = false <-> - labelIn_8 lambda ls1 a = false /\ - labelIn_8 lambda ls2 a = false. - - induction ls1; intros; simpl in *. - destruct a. - intuition. - destruct a. - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) a0 p). - intuition. - discriminate. - eapply IHls1; intuition. - Qed. - - Theorem labelIn_8_false_firstn : - forall Z ls x n, - @labelIn_8 lambda Z (flatten ls) x = false -> - labelIn_8 lambda (flatten (firstn n ls)) x = false. - - induction ls; intuition; simpl in *. - rewrite firstn_nil. - simpl. - trivial. - - destruct n; simpl in *. - trivial. - eapply labelIn_8_app_false in H0. - intuition. - eapply labelIn_8_app_false. - intuition. - - Qed. - - specialize (H O 1%nat x). - simpl in *. - - eapply labelIn_8_false_firstn. - eapply H. - omega. - intuition. - - intuition. - specialize (H O 1%nat x). - simpl in *. - intuition. - - Qed. - - - rewrite (@labelCollision_7_1_skipn _ (length x1)). - rewrite skipn_app. - rewrite firstn_app; intuition. - rewrite firstn_eq_all_gen; intuition. - - rewrite labelCollision_7_app. - f_equal. - eapply IHls1. - - eapply list_pred_app_both_if in H. - intuition. - intuition. - intuition. - eapply H0. - destruct (lt_dec n (length x1)). - rewrite app_nth1; intuition. - rewrite nth_overflow in H2. - simpl in *. - intuition. - omega. - - unfold labelCollision_7_1. - simpl. - rewrite app_nil_r. - rewrite orb_false_r. - - Theorem bucketLabelCollision_7_8_eq : - forall Z ls1 ls2, - list_pred (fun x y => fst x = snd (fst y)) ls1 ls2 -> - (exists b, - (forall a, In a ls2 -> fst (fst a) = b)) -> - bucketLabelCollision_7 ls1 = - @labelCollision_8 lambda Z ls2. - - induction 1; intuition; simpl in *. - destruct a1. - destruct a2. - simpl in *. - subst. - f_equal. - - Theorem labelIn_7_8_eq : - forall Z ls1 ls2 q p, - list_pred (fun x y => fst x = snd (fst y)) ls1 ls2 -> - (forall a, In a ls2 -> fst (fst a) = p) -> - labelIn_7 ls1 q = - @labelIn_8 lambda Z ls2 (p, q). - - induction 1; intuition; simpl in *. - destruct a1. - destruct a2. - simpl in *. - subst. - destruct p1. - simpl in *. - unfold eqbPair. - simpl. - - case_eq (eqb p n); intuition. - simpl. - - destruct (eqbBvector q b); - trivial. - eapply IHlist_pred. - intuition. - - specialize (H1 (n, b, z)). - intuition. - simpl in *. - subst. - rewrite eqb_refl in H. - discriminate. - - Qed. - - destruct p0. - destruct H1. - assert (fst (fst (n, b, z)) = x). - eapply H. - intuition. - simpl in *. - subst. - - eapply labelIn_7_8_eq . - trivial. - intuition. - - eapply IHlist_pred. - destruct H1. - econstructor. - intuition. - - Qed. - - eapply list_pred_app_both_if in H. - intuition. - inversion H4; clear H4; subst. - eapply bucketLabelCollision_7_8_eq. - eapply list_pred_impl. - eapply H7. - intuition. - econstructor. - intuition. - eapply H0. - rewrite app_nth2. - rewrite minus_diag. - simpl. - trivial. - intuition. - intuition. - - intuition. - assert (fst b = j). - - Theorem labelIn_8_exists : - forall Z ls x, - @labelIn_8 lambda Z ls x = true -> - exists z, - In (x, z) ls. - - induction ls; intuition; simpl in *. - discriminate. - - case_eq (eqbPair nat_EqDec (Bvector_EqDec lambda) x (a, b0)); intuition. - rewrite H1 in H0. - unfold eqbPair in *. - rewrite andb_true_iff in H1. - intuition. - rewrite eqb_leibniz in H2. - rewrite eqb_leibniz in H3. - destruct x. - simpl in *; subst. - econstructor. - intuition. - - rewrite H1 in H0. - edestruct IHls. - eauto. - econstructor. - right. - eauto. - Qed. - - eapply labelIn_8_exists in H4. - destruct H4. - specialize (H0 j (b, x2)). - eapply H0. - trivial. - subst. - - subst. - - Theorem labelIn_8_all_false : - forall Z ls b, - (forall x y, In (x, y) ls -> x <> b) -> - @labelIn_8 lambda Z ls b = false. - - induction ls; intuition; simpl in *. - case_eq (eqbPair nat_EqDec (Bvector_EqDec lambda) b1 (a, b0)); intuition. - unfold eqbPair in *. - destruct b1. - simpl in *. - apply andb_true_iff in H1. - intuition. - rewrite eqb_leibniz in H2. - apply eqbBvector_sound in H3. - subst. - exfalso. - eapply H0. - left. - reflexivity. - trivial. - - eapply IHls; intuition. - subst. - eapply H0. - right. - eauto. - trivial. - - Qed. - - eapply labelIn_8_all_false . - intuition. - subst. - eapply in_flatten in H6. - simpl in *. - destruct H6. - intuition. - - eapply nth_In_exists in H7. - destruct H7. - - rewrite nth_skipn_eq in H6. - rewrite <- H6 in H8. - clear H6. - - specialize (H0 (x3 + j')%nat (a, b0, y)). - simpl in H0. - assert (a = (x3 + j')%nat). - eapply H0. - eauto. - subst. - omega. - Qed. - - eapply evalDist_ret_eq. - f_equal. - eapply labelCollision_7_1_eq. - - Theorem ESPADA_TSetSetup_tLoop_Corr_7_1_pred : - forall t1 t2, - list_pred - (fun a b6 => - list_pred - (fun x0 y => - fst x0 = snd (fst y) /\ snd x0 = snd y) a b6) - t1 t2 -> - forall p, - list_pred - (fun a b6 => - list_pred - (fun x0 y => - fst x0 = snd (fst y) /\ snd x0 = snd y) a b6) - (ESPADA_TSetSetup_tLoop_Corr_7 t1 p) - (ESPADA_TSetSetup_tLoop_Corr_7_1 t2 p). - - intuition. - - unfold ESPADA_TSetSetup_tLoop_Corr_7, ESPADA_TSetSetup_tLoop_Corr_7_1. - comp_simp. - - eapply list_pred_listReplace. - intuition. - - eapply list_pred_app_both. - - eapply list_pred_nth. - eapply list_pred_impl. - eapply H. - intuition. - econstructor. - - econstructor. - unfold fst, snd. - intuition. - econstructor. - econstructor. - Qed. - - Theorem ESPADA_TSetSetup_Corr_7_1_pred : - forall ls t1 t2, - list_pred - (fun a b => - list_pred - (fun x0 y => - fst x0 = snd (fst y) /\ snd x0 = snd y) - a b) - t1 t2 -> - list_pred - (fun a b => - list_pred - (fun x0 y => - fst x0 = snd (fst y) /\ snd x0 = snd y) - a b) - (fold_left (@ESPADA_TSetSetup_tLoop_Corr_7 lambda) ls t1) - (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls t2). - - induction ls; intuition. - - Theorem fold_left_cons : - forall (A B : Type)(ls : list A)(a : A)(b : B) f, - fold_left f (a :: ls) b = fold_left f ls (f b a). - - intuition. - Qed. - - rewrite fold_left_cons. - rewrite fold_left_cons. - eapply IHls. - - eapply ESPADA_TSetSetup_tLoop_Corr_7_1_pred. - trivial. - - Qed. - - eapply ESPADA_TSetSetup_Corr_7_1_pred . - econstructor. - - intuition. - - Theorem ESPADA_TSetSetup_tLoop_Corr_7_1_bucket_label : - forall n t a b, - (forall a n, In a (nth n t nil) -> fst (fst a) = n) -> - In a (nth n (ESPADA_TSetSetup_tLoop_Corr_7_1 t b) nil) -> - fst (fst a) = n. - - intuition. - - unfold ESPADA_TSetSetup_tLoop_Corr_7_1 in *. - destruct b. - repeat (destruct p). - unfold setLet in *. - destruct (eq_nat_dec n n0). - subst. - rewrite nth_listReplace_eq in H0. - - eapply in_app_or in H0. - intuition. - simpl in *; intuition. - inversion H0; trivial. - rewrite nth_listReplace_ne in H0; intuition. - Qed. - - Theorem ESPADA_TSetSetup_Corr_7_1_bucket_label : - forall ls a n t, - (forall a n, In a (nth n t nil) -> fst (fst a) = n) -> - In a (nth n (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls t) nil) -> - fst (fst a) = n. - - induction ls; intros. - intuition. - rewrite fold_left_cons in H0. - eapply (IHls _ _ (ESPADA_TSetSetup_tLoop_Corr_7_1 t a)). - intros. - eapply ESPADA_TSetSetup_tLoop_Corr_7_1_bucket_label; eauto. - - trivial. - - Qed. - - eapply ESPADA_TSetSetup_Corr_7_1_bucket_label. - intuition. - rewrite nth_nil in H2. - simpl in *. - intuition. - eauto. - - f_equal. - f_equal. - - eapply map_ext_in. - intuition. - - Theorem ESPADA_TSetRetrieve_tLoop_Corr_7_1_equiv : - forall t1 t2, - list_pred - (fun a b => - list_pred - (fun x y => - fst x = snd (fst y) /\ snd x = snd y) a b) - t1 t2 -> - (forall p n, In p (nth n t2 nil) -> fst (fst p) = n) -> - forall i z, - (ESPADA_TSetRetrieve_tLoop_Corr_7 F) t1 z i = ESPADA_TSetRetrieve_tLoop_Corr_7_1 t2 z i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_7, ESPADA_TSetRetrieve_tLoop_Corr_7_1. - remember (F z i). - comp_simp. - - Theorem ESPADA_TSetRetrieve_Corr_7_1_arrayLookup_equiv : - forall (t1 : TSetCorr_7 lambda) (t2 : TSetCorr_7_1), - list_pred - (fun a b => - list_pred - (fun x y => - fst x = snd (fst y) /\ snd x = snd y) a b) - t1 t2 -> - forall n1 n2 z, - (forall p, In p (nth n1 t2 nil) -> fst (fst p) = n2) -> - arrayLookup _ (nth n1 t1 nil) z = - arrayLookup _ (nth n1 t2 nil) (n2, z). - - induction 1; intuition; simpl in *. - destruct n1; - simpl; - trivial. - - destruct n1; simpl. - - Theorem arrayLookup_pair_const_eq : - forall a1 a2, - list_pred - (fun - (x : Bvector lambda * - (Blist * Bvector lambda * nat * nat * Bvector lambda * nat * - Bvector (S lambda) * Bvector (S lambda))) - (y : nat * Bvector lambda * - (Blist * Bvector lambda * nat * nat * Bvector lambda * nat * - Bvector (S lambda) * Bvector (S lambda))) => - fst x = snd (fst y) /\ snd x = snd y) a1 a2 -> - forall c z, - (forall x, In x a2 -> fst (fst x) = c) -> - arrayLookup _ a1 z = - arrayLookup _ a2 (c, z). - - induction 1; intuition; simpl in *. - destruct a1. destruct a2. - unfold eqbPair. - simpl in *. - destruct p0. - simpl in *. - subst. - case_eq (eqb c n); intuition. - simpl. - case_eq (eqbBvector z b0); intuition. - specialize (H (n, b0, p1)). - intuition. - subst. - simpl in *. - rewrite eqb_refl in H1. - discriminate. - - Qed. - - eapply arrayLookup_pair_const_eq; - intuition. - eapply IHlist_pred. - intuition. - - Qed. - - erewrite ESPADA_TSetRetrieve_Corr_7_1_arrayLookup_equiv; intuition. - Qed. - - - Theorem ESPADA_TSetRetrieve_h_Corr_7_1_equiv : - forall n t1 t2, - list_pred - (fun a b => - list_pred - (fun x y => - fst x = snd (fst y) /\ snd x = snd y) a b) - t1 t2 -> - (forall n p, In p (nth n t2 nil) -> fst (fst p) = n) -> - forall i z, - ESPADA_TSetRetrieve_h_Corr_7 F t1 z i n = ESPADA_TSetRetrieve_h_Corr_7_1 t2 z i n. - - induction n; intuition; simpl in *. - case_eq (ESPADA_TSetRetrieve_tLoop_Corr_7_1 t2 z i); intuition; - erewrite <- ESPADA_TSetRetrieve_tLoop_Corr_7_1_equiv in H2. - erewrite H2. - comp_simp. - f_equal. - destruct b; trivial. - eapply IHn; intuition. - eapply H. - intuition. - - rewrite H2. - trivial. - eapply H. - intuition. - - Qed. - - Theorem ESPADA_TSetRetrieve_Corr_7_1_equiv : - forall t1 t2, - list_pred - (fun a b => - list_pred - (fun x y => - fst x = snd (fst y) /\ snd x = snd y) a b) - t1 t2 -> - (forall n p, In p (nth n t2 nil) -> fst (fst p) = n) -> - forall z, - ESPADA_TSetRetrieve_Corr_7 maxMatches F t1 z = ESPADA_TSetRetrieve_Corr_7_1 t2 z. - - intuition. - eapply ESPADA_TSetRetrieve_h_Corr_7_1_equiv ; intuition. - - Qed. - - eapply ESPADA_TSetRetrieve_Corr_7_1_equiv . - eapply ESPADA_TSetSetup_Corr_7_1_pred . - econstructor. - intuition. - - eapply ESPADA_TSetSetup_Corr_7_1_bucket_label. - intuition. - rewrite nth_nil in H2. - simpl in *. - intuition. - eauto. - Qed. - - - Theorem ESPADA_TSetCorr_G7_1_G8_equiv : - Pr[ESPADA_TSetCorr_G7_1] == Pr[ESPADA_TSetCorr_G8 maxMatches F A1]. - - intuition. - unfold ESPADA_TSetCorr_G7_1, ESPADA_TSetCorr_G8. - comp_simp. - comp_skip. - comp_simp. - comp_skip. - eapply evalDist_ret_eq. - - f_equal. - - Theorem labelCollision_8_perm_eq : - forall Z ls1 ls2, - Permutation ls1 ls2 -> - @labelCollision_8 lambda Z ls1 = - labelCollision_8 lambda ls2. - - induction 1; intuition; simpl in *. - - Theorem labelIn_8_perm_eq : - forall Z ls1 ls2, - Permutation ls1 ls2 -> - forall p, - @labelIn_8 lambda Z ls1 p = labelIn_8 lambda ls2 p. - - induction 1; intuition; simpl in *. - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) (a, b1) (a0, b0)); intuition. - - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) (a0, b3) (a, b2)). - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) (a0, b3) (a1, b1)); intuition. - trivial. - - rewrite IHPermutation1. - eauto. - - Qed. - - f_equal. - eapply labelIn_8_perm_eq; intuition. - - trivial. - - rewrite eqbPair_symm. - destruct (eqbPair nat_EqDec (Bvector_EqDec lambda) (a1, b1) (a, b2)); intuition. - - destruct (labelIn_8 lambda l (a1, b1)); destruct ((labelIn_8 lambda l (a, b2) )); intuition. - - rewrite IHPermutation1. - trivial. - - Qed. - - eapply labelCollision_8_perm_eq. - - - Theorem setup_tLoop_Corr_7_1_8_eq : - forall t1 t2 a, - Permutation (flatten t1) t2 -> - Permutation (flatten (ESPADA_TSetSetup_tLoop_Corr_7_1 t1 a)) - (t2 ++ (ESPADA_TSetSetup_tLoop_Corr_8 lambda a) :: nil). - - intuition. - unfold ESPADA_TSetSetup_tLoop_Corr_7_1, ESPADA_TSetSetup_tLoop_Corr_8. - comp_simp. - - eapply perm_trans. - - eapply perm_flatten_listReplace. - eauto. - - eapply Permutation_cons_append. - Qed. - - - Theorem setup_Corr_7_1_8_eq_h : - forall ls t1 t2, - Permutation (flatten t1) t2 -> - Permutation (flatten (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls t1)) - (t2 ++ (map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls)). - - induction ls; intuition. - simpl. - rewrite app_nil_r. - trivial. - - repeat rewrite fold_left_cons. - - rewrite map_cons. - rewrite app_cons_eq. - eapply IHls. - - eapply setup_tLoop_Corr_7_1_8_eq. - trivial. - Qed. - - Theorem setup_Corr_7_1_8_eq : - forall ls, - Permutation (flatten (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls nil)) - ((map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls)). - - intuition. - eapply perm_trans. - eapply setup_Corr_7_1_8_eq_h. - simpl. - eapply Permutation_refl. - simpl. - eapply Permutation_refl. - - Qed. - - eapply setup_Corr_7_1_8_eq. - - f_equal. - f_equal. - - eapply map_ext_in. - intuition. - - - Theorem retrieve_tLoop_Corr_7_1_8_eq : - forall i t1 t2 x, - (forall b l, arrayLookup _ (nth b t1 nil) (b, l) = arrayLookup _ t2 (b, l)) -> - ESPADA_TSetRetrieve_tLoop_Corr_7_1 t1 x i = - ESPADA_TSetRetrieve_tLoop_Corr_8 F t2 x i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_7_1, ESPADA_TSetRetrieve_tLoop_Corr_8. - remember (F x i). - comp_simp. - - rewrite H. - destruct (arrayLookup (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) t2 (n, b0)); intuition. - - Qed. - - Theorem retrieve_h_Corr_7_1_8_eq : - forall fuel i t1 t2 x, - (forall b l, arrayLookup _ (nth b t1 nil) (b, l) = arrayLookup _ t2 (b, l)) -> - ESPADA_TSetRetrieve_h_Corr_7_1 t1 x i fuel = - ESPADA_TSetRetrieve_h_Corr_8 F t2 x i fuel. - - induction fuel; intuition; simpl in *. - case_eq (ESPADA_TSetRetrieve_tLoop_Corr_7_1 t1 x i); intuition; - erewrite retrieve_tLoop_Corr_7_1_8_eq in H1. - rewrite H1. - comp_simp. - f_equal. - destruct b; trivial. - eapply IHfuel. - intuition. - intuition. - rewrite H1. - trivial. - intuition. - - Qed. - - - Theorem retrieve_Corr_7_1_8_eq : - forall t1 t2 x, - (forall b l, arrayLookup _ (nth b t1 nil) (b, l) = arrayLookup _ t2 (b, l)) -> - ESPADA_TSetRetrieve_Corr_7_1 t1 x = - ESPADA_TSetRetrieve_Corr_8 maxMatches F t2 x. - - intuition. - eapply retrieve_h_Corr_7_1_8_eq . - intuition. - - Qed. - - - eapply retrieve_Corr_7_1_8_eq. - intuition. - - Theorem setup_tLoop_Corr_7_1_8_arrayLookup_eq : - forall t1 t2 a, - (forall b l, arrayLookup _ (nth b t1 nil) (b, l) = arrayLookup _ t2 (b, l)) -> - (forall b l, arrayLookup _ (nth b (ESPADA_TSetSetup_tLoop_Corr_7_1 t1 a) nil) (b, l) = - arrayLookup _ (t2 ++ ((@ESPADA_TSetSetup_tLoop_Corr_8 lambda) a :: nil)) (b, l)). - - intuition. - unfold ESPADA_TSetSetup_tLoop_Corr_7_1, ESPADA_TSetSetup_tLoop_Corr_8. - comp_simp. - destruct (eq_nat_dec b6 b1). - subst. - rewrite nth_listReplace_eq. - - case_eq (arrayLookup _ (nth b1 t1 nil) (b1, l)); intuition. - - erewrite arrayLookup_app_Some; eauto. - rewrite H in H0. - erewrite arrayLookup_app_Some; eauto. - - erewrite arrayLookup_app_None; eauto. - rewrite H in H0. - erewrite arrayLookup_app_None; eauto. - - rewrite nth_listReplace_ne; intuition. - - rewrite H. - rewrite arrayLookup_app_ne; intuition. - pairInv. - intuition. - Qed. - - - Theorem setup_Corr_7_1_8_arrayLookup_eq_h : - forall ls t1 t2, - (forall b l, arrayLookup _ (nth b t1 nil) (b, l) = arrayLookup _ t2 (b, l)) -> - (forall b l, arrayLookup _ (nth b (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls t1) nil) (b, l) = - arrayLookup _ (t2 ++ (map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls)) (b, l)). - - induction ls; intuition. - simpl in *. - rewrite app_nil_r. - trivial. - - repeat rewrite fold_left_cons. - rewrite map_cons. - rewrite app_cons_eq. - eapply IHls. - eapply setup_tLoop_Corr_7_1_8_arrayLookup_eq. - trivial. - Qed. - - Theorem setup_Corr_7_1_8_arrayLookup_eq : - forall ls, - (forall b l, arrayLookup _ (nth b (fold_left ESPADA_TSetSetup_tLoop_Corr_7_1 ls nil) nil) (b, l) = - arrayLookup _ (map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls) (b, l)). - - intuition. - erewrite setup_Corr_7_1_8_arrayLookup_eq_h. - rewrite app_nil_l. - trivial. - intuition. - simpl. - destruct b0; trivial. - Qed. - - eapply setup_Corr_7_1_8_arrayLookup_eq. - - Qed. - - Theorem ESPADA_TSetCorr_G7_G8_equiv : - Pr[ESPADA_TSetCorr_G7 maxMatches F A1] == Pr[ESPADA_TSetCorr_G8 maxMatches F A1]. - - intuition. - - rewrite ESPADA_TSetCorr_G7_1_equiv. - eapply ESPADA_TSetCorr_G7_1_G8_equiv . - - Qed. - - (* check results for all queries---this simplifies things a bit. *) - Definition ESPADA_TSetCorr_G8_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allTags) (ESPADA_TSetRetrieve_Corr_8 maxMatches F tSet x)) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret (labelCollision_8 _ (tSet) || bad). - - Theorem ESPADA_TSetCorr_G8_8_1_equiv: - Pr[ESPADA_TSetCorr_G8 maxMatches F A1] <= Pr[ESPADA_TSetCorr_G8_1]. - - intuition. - unfold ESPADA_TSetCorr_G8, ESPADA_TSetCorr_G8_1. - comp_simp. - comp_skip. - comp_simp. - comp_skip. - eapply comp_spec_impl_le. - eapply comp_spec_ret. - intuition. - eapply orb_true_iff in H1. - intuition. - eapply orb_true_iff. - right. - - eapply negb_true_iff in H2. - eapply eqb_false_iff in H2. - - eapply negb_true_iff. - eapply eqb_false_iff. - intuition. - eapply H2. - - rewrite <- (@firstn_skipn _ (length (removeDups (list_EqDec bool_EqDec) l)) x) in H1 at 1. - unfold combineKeywords in H1. - repeat rewrite map_app in H1. - - Show. - eapply app_eq_inv in H1. - intuition. - clear H4. - unfold combineKeywords. - repeat rewrite map_app. - eapply H3. - - repeat rewrite map_length. - rewrite firstn_length. - eapply min_l. - eapply compMap_length in H0. - rewrite H0. - unfold combineKeywords. - rewrite app_length. - omega. - - Qed. - - Theorem ESPADA_TSetCorr_G8_1_9_equiv: - Pr[ESPADA_TSetCorr_G8_1] <= Pr[ESPADA_TSetCorr_G9 maxMatches F A1]. - - intuition. - unfold ESPADA_TSetCorr_G8_1, ESPADA_TSetCorr_G9. - comp_simp. - comp_skip; comp_simp. - comp_skip. - eapply comp_spec_impl_le. - eapply comp_spec_ret. - intuition. - eapply orb_true_iff in H1. - intuition. - - eapply orb_true_iff. - left. - - Theorem labelIn_8_In : - forall Z ls p, - @labelIn_8 lambda Z ls p = true <-> - In p (fst (split ls)). - - induction ls; intuition; simpl in *. - discriminate. - remember (split ls) as z. - destruct z. - simpl in *. - case_eq (eqbPair nat_EqDec (Bvector_EqDec lambda) (a0, b1) (a, b0)); intuition. - unfold eqbPair in *. - simpl in *. - eapply andb_true_iff in H1. - intuition. - rewrite eqb_leibniz in H2. - eapply eqbBvector_sound in H3. - subst. - intuition. - - rewrite H1 in H0. - right. - eapply IHls. - trivial. - - remember (split ls) as z. - destruct z. - simpl in *. - intuition. - pairInv. - unfold eqbPair; simpl. - rewrite eqb_refl. - rewrite eqbBvector_complete. - simpl. - trivial. - - destruct ( eqbPair nat_EqDec (Bvector_EqDec lambda) (a0, b1) (a, b0)); trivial. - eapply IHls. - trivial. - - Qed. - - Theorem labelCollision_8_false_NoDup : - forall Z ls, - @labelCollision_8 lambda Z ls = false <-> - NoDup (fst (split ls)). - - induction ls; intuition; simpl in *. - econstructor. - remember (split ls) as z. - destruct z. - simpl in *. - eapply orb_false_iff in H2. - intuition. - - econstructor. - intuition. - specialize (@labelIn_8_In _ ls (a, b0)); intuition. - rewrite <- Heqz in H7. - simpl in *. - intuition. - rewrite H5 in H3. - discriminate. - - trivial. - - remember (split ls) as z. - destruct z. - simpl in *. - inversion H2; clear H2; subst. - intuition. - case_eq (labelIn_8 _ ls (a, b0)); intuition. - exfalso. - eapply H5. - specialize (@labelIn_8_In _ ls (a, b0)); intuition. - rewrite <- Heqz in H3. - simpl in *. - trivial. - - Qed. - - case_eq (labelCollision_9 _ maxMatches F x); intuition. - eapply hasDups_false_NoDup in H1. - - match goal with - | [H : labelCollision_8 _ ?a = true |- _ ] => assert (labelCollision_8 _ a = false) - end. - - Fixpoint labelIn_9_1 (ls : list (Blist * Bvector lambda * nat * nat * Bvector lambda * nat * Bvector lambda * Bvector (S lambda))) p := - match ls with - | nil => false - | (_, _, _, _, _, b', l', _) :: ls' => - eqb p (b', l') || labelIn_9_1 ls' p - end. - - Fixpoint labelCollision_9_1 (ls : list (Blist * Bvector lambda * nat * nat * Bvector lambda * nat * Bvector lambda * Bvector (S lambda))) := - match ls with - | nil => false - | (_, _, _, _, _, b, l, _) :: ls' => - labelIn_9_1 ls' (b, l) || labelCollision_9_1 ls' - end. - - Theorem labelIn_8_9_1_equiv : - forall ls b l , - labelIn_8 _ (map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls) (b, l) = labelIn_9_1 ls (b, l). - - induction ls; intuition; simpl in *. - case_eq (eqbPair nat_EqDec (Bvector_EqDec lambda) (b6, l) (b1, b0)); intuition. - - simpl. - eapply IHls; eauto. - Qed. - - Theorem labelCollision_8_9_1_equiv : - forall ls, - labelCollision_8 _ (map (@ESPADA_TSetSetup_tLoop_Corr_8 lambda) ls) = labelCollision_9_1 ls. - - induction ls; intuition; simpl in *. - f_equal. - eapply labelIn_8_9_1_equiv . - eauto. - - Qed. - - rewrite labelCollision_8_9_1_equiv. - - Theorem labelIn9_1_In_eq : - forall ls1 (ls2 : list (nat * Bvector lambda)) p, - list_pred (fun x y => let '(_, _, _, _, _, b, l, _) := x in (b, l) = y) ls1 ls2 -> - labelIn_9_1 ls1 p = - if (in_dec (EqDec_dec _) p ls2) then true else false. - - induction 1; intuition; simpl in *. - destruct a1. - repeat (destruct p0). - - destruct (EqDec_dec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) a2 p). - subst. - subst. - unfold eqbPair; simpl. - repeat rewrite eqb_refl. - rewrite eqbBvector_complete. - simpl. - trivial. - subst. - case_eq ( eqbPair nat_EqDec (Bvector_EqDec lambda) p (n, b0)); intuition. - unfold eqbPair in *; simpl in *. - eapply andb_true_iff in H. - intuition. - rewrite - eqb_leibniz in H1. - apply eqbBvector_sound in H3. - subst. - exfalso; destruct p; simpl in *; intuition. - - simpl. - rewrite IHlist_pred. - destruct ( in_dec (EqDec_dec (pair_EqDec nat_EqDec (Bvector_EqDec lambda))) p ls2); intuition. - - Qed. - - Theorem labelCollision_9_1_hasDups_eq : - forall ls1 (ls2 : list (nat * Bvector lambda)), - list_pred (fun x y => let '(_, _, _, _, _, b, l, _) := x in (b, l) = y) ls1 ls2 -> - labelCollision_9_1 ls1 = - hasDups _ ls2. - - induction 1; intuition; simpl in *. - destruct a1. - repeat (destruct p). - subst. - rewrite (@labelIn9_1_In_eq _ ls2). - destruct (in_dec (EqDec_dec (pair_EqDec nat_EqDec (Bvector_EqDec lambda))) - (n, b0) ls2); intuition. - intuition. - - Qed. - - rewrite (@labelCollision_9_1_hasDups_eq _ - (flatten - (foreach (p in - (zip x - (map (fun z => length (arrayLookupList (list_EqDec bool_EqDec) t z)) - (combineKeywords (removeDups (list_EqDec bool_EqDec) l) - (toW t))) - ) - ) - [tag, n]<-2 p; - numberedMap - (fun (i len : nat) (s_i : nat) => - [b, bigL, bigK]<-3 F tag i; (b, bigL)) - (allNatsLt n))) - ). - - eapply hasDups_false_NoDup. - - Definition getAllLabels_8_1 tags := - map (fun t => map (fun i => fst (F t i)) (allNatsLt (S maxMatches))) tags. - - - Theorem getAllLabels_8_1_equiv : - forall ls, - getAllLabels _ maxMatches F ls = flatten (getAllLabels_8_1 ls). - - induction ls; intuition; simpl in *. - f_equal. - eauto. - - Qed. - - rewrite getAllLabels_8_1_equiv in H1. - - eapply (@NoDup_flatten_subset _ (getAllLabels_8_1 x)). - unfold getAllLabels_8_1. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_zip_r. - - assert (list_pred (fun a b => b <= maxMatches)%nat x - (foreach (z - in combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) - length (arrayLookupList (list_EqDec bool_EqDec) t z))). - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_I. - eapply compMap_length. - eauto. - intuition. - - eapply arrayLookupList_pred. - intuition. - eapply maxMatches_correct. - eauto. - trivial. - simpl; intuition. - - eapply H3. - eapply list_pred_eq. - eapply list_pred_I. - rewrite map_length. - symmetry. - eapply compMap_length. - eauto. - - intuition. - intuition. - subst. - exists (snd b). - destruct b. - unfold numberedMap. - simpl. - - rewrite firstn_map. - rewrite allNatsLt_length. - - eapply map_eq. - rewrite (@firstn_allNatsLt _ (S maxMatches)). - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_eq. - eapply list_pred_eq. - eapply list_pred_eq. - intuition. - simpl in *; subst. - destruct (F b b1). - destruct p. - trivial. - simpl in *. - omega. - - trivial. - - eapply list_pred_flatten_both. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_zip_l. - eapply list_pred_I. - symmetry. - eapply compMap_length; eauto. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_I. - rewrite map_length. - eapply compMap_length; eauto. - - eapply list_pred_zip_l. - eapply list_pred_I. - symmetry. - eapply compMap_length; eauto. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_I. - eapply compMap_length; eauto. - eapply list_pred_I. - eapply compMap_length; eauto. - - eapply list_pred_map_l. - eapply list_pred_eq. - - eapply list_pred_zip_r. - eapply list_pred_I. - rewrite map_length. - eapply compMap_length; eauto. - eapply list_pred_eq. - - eapply list_pred_I. - symmetry. - rewrite map_length. - eapply compMap_length; eauto. - - eapply list_pred_zip_r. - eapply list_pred_I. - rewrite map_length. - eapply compMap_length; eauto. - - eapply list_pred_I. - rewrite map_length. - eapply compMap_length; eauto. - - eapply list_pred_map_both. - eapply list_pred_eq. - - intuition. - destruct H6. - destruct H13. - destruct H13. - destruct H11. - destruct H16. - intuition; simpl in *; subst. - destruct b1; simpl in *; subst. - - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - - rewrite allNatsLt_length. - - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - apply allNatsLt_length. - - eapply list_pred_zip_r. - eapply list_pred_eq. - rewrite H21. - eapply list_pred_eq. - rewrite H21. - eapply list_pred_eq. - - eapply list_pred_zip_r. - eapply list_pred_eq. - rewrite H21. - eapply list_pred_I. - apply allNatsLt_length. - rewrite H21. - eapply list_pred_I. - apply allNatsLt_length. - - intuition. - destruct b1. - simpl in *; subst. - remember (F b a0) as z. - destruct z. - destruct p. - trivial. - - rewrite H3 in H2. - discriminate. - Qed. - - Theorem ESPADA_TSetCorr_G9_G10_equiv : - Pr[ESPADA_TSetCorr_G9 maxMatches F A1] == Pr[ESPADA_TSetCorr_G10 maxMatches F A1]. - - intuition. - unfold ESPADA_TSetCorr_G9, ESPADA_TSetCorr_G10. - comp_simp. - comp_skip; comp_simp. - comp_skip. - - eapply evalDist_ret_eq. - eapply orb_same_eq_if. - intuition. - - f_equal. - f_equal. - eapply map_ext_in. - intuition. - - - Theorem tSetRetrieve_tLoop_Corr_8_10_eq : - forall (ls : TSetCorr_8 lambda) a i, - (forall n b0 b4 a' i' n1 b3 n0 b2 b1, - In (n, b0, (b4, a', i', n1, b3, n0, b2, b1)) ls -> - (n, b0) = fst (F a' i') /\ b2 = (snd (F a' i')) /\ b1 = (Vector.cons _ (if (eq_nat_dec (S i') n1) then false else true) _ b3) xor b2 /\ - (fst (F a' i') = fst (F a i) -> a = a' /\ i = i')) -> - ESPADA_TSetRetrieve_tLoop_Corr_8 F ls a i = - ESPADA_TSetRetrieve_tLoop_Corr_10 F ls a i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_8, ESPADA_TSetRetrieve_tLoop_Corr_10. - remember (F a i). - comp_simp. - match goal with - | [|- match ?a with | Some _ => _ | None => _ end = _ ] => case_eq a; intuition - end. - - repeat (destruct p). - - eapply arrayLookup_Some_impl_In in H1. - eapply H0 in H1. - intuition. - subst. - simpl in H5. - symmetry in H2. - intuition. - subst. - - rewrite <- Heqp. - unfold snd. - rewrite BVxor_assoc. - rewrite BVxor_same_id. - rewrite BVxor_id_r. - trivial. - - Qed. - - Theorem tSetRetrieve_h_Corr_8_10_eq : - forall fuel i ls a, - ( i + fuel <= maxMatches)%nat -> - (forall n b0 b4 a' i' n1 b3 n0 b2 b1, - forall i, - (i <= maxMatches)%nat -> - In (n, b0, (b4, a', i', n1, b3, n0, b2, b1)) ls -> - (n, b0) = fst (F a' i') /\ b2 = (snd (F a' i')) /\ b1 = (Vector.cons _ (if (eq_nat_dec (S i') n1) then false else true) _ b3) xor b2 /\ - (fst (F a' i') = fst (F a i) -> a = a' /\ i = i')) -> - ESPADA_TSetRetrieve_h_Corr_8 F ls a i fuel = - ESPADA_TSetRetrieve_h_Corr_10 F ls a i fuel. - - induction fuel; intuition. - simpl. - erewrite tSetRetrieve_tLoop_Corr_8_10_eq . - remember (ESPADA_TSetRetrieve_tLoop_Corr_10 F ls a i). - destruct o. - destruct p. - f_equal. - destruct b0. - - eapply IHfuel. - omega. - - eauto. - trivial. - trivial. - - intros. - eapply H0. - omega. - eauto. - - Qed. - - - Theorem tSetRetrieve_Corr_8_10_eq : - forall ls a, - (forall n b0 b4 a' i' n1 b3 n0 b2 b1, - forall i, - (i <= maxMatches)%nat -> - In (n, b0, (b4, a', i', n1, b3, n0, b2, b1)) ls -> - (n, b0) = fst (F a' i') /\ b2 = (snd (F a' i')) /\ b1 = (Vector.cons _ (if (eq_nat_dec (S i') n1) then false else true) _ b3) xor b2 /\ - (fst (F a' i') = fst (F a i) -> a = a' /\ i = i')) -> - ESPADA_TSetRetrieve_Corr_8 maxMatches F ls a = - ESPADA_TSetRetrieve_Corr_10 maxMatches F ls a. - - intuition. - - eapply tSetRetrieve_h_Corr_8_10_eq; intuition. - - Qed. - - eapply tSetRetrieve_Corr_8_10_eq; intros. - - - Theorem labelCollision_F_1_to_1 : - forall ls a1 a2 i1 i2, - NoDup (getAllLabels _ maxMatches F ls) -> - In a1 ls -> - In a2 ls -> - (i1 <= maxMatches)%nat -> - (i2 <= maxMatches)%nat -> - fst (F a1 i1) = fst (F a2 i2) -> - a1 = a2 /\ i1 = i2. - - induction ls; intros. - simpl in *. - intuition. - - simpl in *. - destruct H0. - subst. - destruct H1. - subst. - - - eapply NoDup_app_l in H. - - eapply NoDup_map in H. - intuition. - eapply H1. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - trivial. - - eapply NoDup_app in H. - destruct H. - destruct H1. - exfalso. - eapply H5. - eapply in_map_iff. - exists (i1). - intuition. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - - assert (In (fst (F a2 i2)) (getAllLabels _ maxMatches F ls)). - - Theorem In_getAllLabels : - forall ls a i, - In a ls -> - (i <= maxMatches)%nat -> - In (fst (F a i)) (getAllLabels _ maxMatches F ls). - - induction ls; intuition; simpl in *. - - intuition; subst. - eapply in_or_app. - left. - eapply in_map_iff. - econstructor. - intuition. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - - Qed. - - eapply In_getAllLabels; intuition. - - eauto. - trivial. - - destruct H1. - subst. - eapply NoDup_app in H. - destruct H. - destruct H1. - exfalso. - eapply H5. - eapply in_map_iff. - exists i2. - intuition. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - - eapply In_getAllLabels. - eapply H0. - eapply H2. - - symmetry. - trivial. - - eapply IHls; intuition. - eapply NoDup_app in H. - intuition. - Qed. - - Ltac intac := - match goal with - | [H: In _ (map _ _) |- _ ] => eapply in_map_iff in H - | [H: exists _, _ |- _] => destruct H - | [H: In _ (flatten _) |- _ ] => eapply in_flatten in H - | [H: In _ (zip _ _) |- _ ] => eapply In_zip in H - | [H: _ /\ _ |- _ ] => destruct H - end. - - repeat (intac; subst). - destruct x2. - destruct p. - unfold numberedMap, ESPADA_TSet_Once_Games.numberedMap in*. - repeat (intac; subst). - destruct x1. - remember (F b5 n2) as z. - destruct z. - destruct p. - unfold ESPADA_TSetSetup_tLoop_Corr_8 in *. - - Ltac pairInv_once' := - match goal with - | [H : (_, _) = (_, _) |- _] => eapply pair_eq_inv in H - | [H : _ /\ _ |- _ ] => destruct H - end. - - Ltac pairInv' := repeat (pairInv_once'; subst). - - pairInv'. - - specialize (@labelCollision_F_1_to_1 x a a' i i'); intuition. - - rewrite <- Heqz. - trivial. - rewrite <- Heqz. - trivial. - - eapply H5; intros; eauto. - eapply hasDups_false_NoDup. - eapply H1. - eapply In_zip in H7. - intuition. - - eapply allNatsLt_lt in H11. - eapply le_trans. - eapply lt_le_weak. - eapply H11. - - eapply arrayLookupList_pred. - intuition. - eapply maxMatches_correct. - eauto. - trivial. - simpl; intuition. - - eapply H5; intros; eauto. - eapply hasDups_false_NoDup. - eapply H1. - eapply In_zip in H7. - intuition. - - eapply allNatsLt_lt in H11. - eapply le_trans. - eapply lt_le_weak. - eapply H11. - - eapply arrayLookupList_pred. - intuition. - eapply maxMatches_correct. - eauto. - trivial. - simpl; intuition. - - Qed. - - Theorem ESPADA_TSetCorr_G8_le_G9: - Pr [ESPADA_TSetCorr_G8 maxMatches F A1] <= Pr [ESPADA_TSetCorr_G9 maxMatches F A1]. - - intuition. - rewrite ESPADA_TSetCorr_G8_8_1_equiv. - eapply ESPADA_TSetCorr_G8_1_9_equiv. - Qed. - - - Theorem TSetRetrieve_tLoop_Corr_10_11_equiv : - forall i t1 t2 x, - list_pred (fun x y => fst x = fst y /\ (snd y) = fst (fst (fst (snd x)))) t2 t1 -> - ESPADA_TSetRetrieve_tLoop_Corr_10 F t2 x i = - ESPADA_TSetRetrieve_tLoop_Corr_11 F t1 x i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_10, ESPADA_TSetRetrieve_tLoop_Corr_11. - remember (F x i). - comp_simp. - - eapply (@list_pred_impl_arrayLookup _ _ _ (fun x y => y = fst (fst (fst x)))) in H. - intuition. - case_eq (arrayLookup (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) t2 (n, b0)); intuition. - repeat (destruct p). - eapply H1 in H. - destruct H. - intuition. - simpl in H4. - subst. - rewrite H3. - trivial. - - rewrite H2; - trivial. - - Qed. - - Theorem TSetRetrieve_h_Corr_10_11_equiv : - forall fuel i t1 t2 x, - list_pred (fun x y => fst x = fst y /\ (snd y) = fst (fst (fst (snd x)))) t2 t1 -> - ESPADA_TSetRetrieve_h_Corr_10 F t2 x i fuel = - ESPADA_TSetRetrieve_h_Corr_11 F t1 x i fuel. - - induction fuel; intuition; simpl in *. - rewrite (@TSetRetrieve_tLoop_Corr_10_11_equiv _ t1). - destruct (ESPADA_TSetRetrieve_tLoop_Corr_11 F t1 x i). - destruct p. - f_equal. - destruct b0. - eapply IHfuel; intuition. - trivial. - trivial. - trivial. - - Qed. - - Theorem TSetRetrieve_Corr_10_11_equiv : - forall t1 t2 x, - list_pred (fun x y => fst x = fst y /\ (snd y) = fst (fst (fst (snd x)))) t2 t1 -> - ESPADA_TSetRetrieve_Corr_10 maxMatches F t2 x = - ESPADA_TSetRetrieve_Corr_11 maxMatches F t1 x. - - eapply TSetRetrieve_h_Corr_10_11_equiv; intuition. - - Qed. - - Theorem ESPADA_TSetCorr_G10_G11_equiv : - Pr[ ESPADA_TSetCorr_G10 maxMatches F A1] == Pr[ ESPADA_TSetCorr_G11 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G10, ESPADA_TSetCorr_G11. - comp_skip. - comp_simp. - comp_skip. - eapply evalDist_ret_eq. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - eapply TSetRetrieve_Corr_10_11_equiv. - - eapply list_pred_map_l'. - eapply list_pred_flatten_both. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - - intuition; subst. - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - - unfold ESPADA_TSetSetup_tLoop_Corr_8. - remember (F b0 a2); intuition. - destruct p. - destruct p. - trivial. - - unfold ESPADA_TSetSetup_tLoop_Corr_8. - remember (F b0 a2); intuition. - destruct p. - destruct p. - trivial. - - Qed. - - - Theorem TSetRetrieve_tLoop_Corr_11_12_equiv : - forall i (t1 : TSetCorr_11 lambda) t2 x, - - list_pred (fun x y => - let '(b1, l1, (q1, z1, n1, c1, s1)) := x in - let '(z2, n2, (q2, c2, s2)) := y in - z1 = z2 /\ n1 = n2 /\ q1 = q2 /\ c1 = c2 /\ s1 = s2 /\ - (b1, l1) = fst (F z1 n1)) t1 t2 -> - NoDup (fst (split t1)) -> - (In (fst (F x i)) (fst (split t1)) -> In (x, i) (fst (split t2))) -> - ESPADA_TSetRetrieve_tLoop_Corr_11 F t1 x i = - ESPADA_TSetRetrieve_tLoop_Corr_12 t2 x i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_11, ESPADA_TSetRetrieve_tLoop_Corr_12. - remember (F x i); intuition. - destruct p. - destruct p. - - case_eq (arrayLookup _ t2 (x, i)); intuition. - apply arrayLookup_Some_impl_In in H3. - eapply list_pred_symm in H. - eapply list_pred_In_exists in H; eauto. - destruct H. - intuition. - destruct x0. - destruct p0. - destruct p1. - repeat (destruct p0). - repeat (destruct p). - intuition; subst. - - rewrite <- Heqp in H10. - simpl in H10. - pairInv. - erewrite arrayLookup_In_NoDup; eauto. - unfold setLet. - trivial. - - case_eq (arrayLookup (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) t1 (n, b0)); intuition. - exfalso. - - - eapply arrayLookup_None_not_In in H3. - intuition. - eapply H1. - simpl. - apply arrayLookup_Some_impl_In in H4. - eapply in_split_l in H4. - trivial. - - Qed. - - Theorem TSetRetrieve_h_Corr_11_12_equiv : - forall fuel i t1 t2 x, - list_pred (fun x y => - let '(b1, l1, (q1, z1, n1, c1, s1)) := x in - let '(z2, n2, (q2, c2, s2)) := y in - z1 = z2 /\ n1 = n2 /\ q1 = q2 /\ c1 = c2 /\ s1 = s2 /\ - (b1, l1) = fst (F z1 n1)) t1 t2 -> - NoDup (fst (split t1)) -> - (i + fuel <= maxMatches)%nat -> - (forall i, - (i <= maxMatches)%nat -> - (In (fst (F x i)) (fst (split t1)) -> In (x, i) (fst (split t2)))) -> - ESPADA_TSetRetrieve_h_Corr_11 F t1 x i fuel = - ESPADA_TSetRetrieve_h_Corr_12 t2 x i fuel. - - induction fuel; intuition; simpl in *. - rewrite (@TSetRetrieve_tLoop_Corr_11_12_equiv _ _ t2). - destruct (ESPADA_TSetRetrieve_tLoop_Corr_12 t2 x i). - destruct p. - f_equal. - destruct b0. - eapply IHfuel; intuition. - trivial. - trivial. - trivial. - trivial. - eapply H2. - omega. - - Qed. - - Theorem TSetRetrieve_Corr_11_12_equiv : - forall t1 t2 x, - list_pred (fun x y => - let '(b1, l1, (q1, z1, n1, c1, s1)) := x in - let '(z2, n2, (q2, c2, s2)) := y in - z1 = z2 /\ n1 = n2 /\ q1 = q2 /\ c1 = c2 /\ s1 = s2 /\ - (b1, l1) = fst (F z1 n1)) t1 t2 -> - NoDup (fst (split t1)) -> - (forall i, - (i <= maxMatches)%nat -> - (In (fst (F x i)) (fst (split t1)) -> In (x, i) (fst (split t2)))) -> - - ESPADA_TSetRetrieve_Corr_11 maxMatches F t1 x = - ESPADA_TSetRetrieve_Corr_12 maxMatches t2 x. - - intuition. - eapply TSetRetrieve_h_Corr_11_12_equiv; intuition. - - Qed. - - Theorem ESPADA_TSetCorr_G11_G12_equiv : - Pr[ESPADA_TSetCorr_G11 maxMatches F A1] == Pr[ESPADA_TSetCorr_G12 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G11, ESPADA_TSetCorr_G12. - - comp_skip; comp_simp. - comp_skip. - eapply evalDist_ret_eq. - case_eq (labelCollision_9 lambda maxMatches F x); intuition. - simpl. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - eapply TSetRetrieve_Corr_11_12_equiv. - - eapply list_pred_flatten_both. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - remember (F b0 a2). - destruct p. - destruct p. - intuition. - rewrite <- Heqp. - trivial. - - rewrite fst_split_flatten_eq. - unfold labelCollision_9 in *. - rewrite getAllLabels_8_1_equiv in H1. - eapply hasDups_false_NoDup in H1. - eapply NoDup_flatten_subset; eauto. - eapply list_pred_map_r'. - eapply list_pred_map_r'. - - - eapply list_pred_impl. - - assert (list_pred (fun a b => a = fst (split (map (F (snd (fst b))) (allNatsLt (S maxMatches)))) /\ length (snd b) <= maxMatches)%nat - (getAllLabels_8_1 x) - (zip - (zip (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) - x) - (foreach (x - in combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) - arrayLookupList (list_EqDec bool_EqDec) t x)) - ). - - eapply list_pred_impl. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_zip_l. - eapply list_pred_I. - symmetry. - eapply compMap_length. - eauto. - eapply list_pred_eq. - eapply list_pred_I. - eapply compMap_length; eauto. - eapply list_pred_zip_l. - eapply list_pred_I. - symmetry. - eapply compMap_length. - eauto. - eapply list_pred_I. - - unfold getAllLabels_8_1. - rewrite map_length. - symmetry. - eapply compMap_length. - eauto. - - Theorem getAllLabels_8_1_list_pred : - forall x, - list_pred - (fun a b => b = fst (split (map (F a) (allNatsLt (S maxMatches))))) - x (getAllLabels_8_1 x). - - induction x; intuition; simpl in *. - econstructor. - econstructor. - symmetry. - eapply fst_split_map_eq. - trivial. - Qed. - - eapply getAllLabels_8_1_list_pred. - eapply list_pred_map_l. - eapply list_pred_I. - unfold getAllLabels_8_1. - rewrite map_length. - symmetry. - eapply compMap_length. - eauto. - intuition. - intuition. - intuition. - destruct H5. - destruct H7. - intuition. - subst. - rewrite H11. - eapply arrayLookupList_pred. - eapply maxMatches_correct. - eauto. - simpl. - intuition. - - eapply H4. - intuition. - intuition. - subst. - destruct b. - destruct p. - simpl in *. - exists (length l0). - unfold ESPADA_TSet_Once_Games.numberedMap. - - rewrite fst_split_map_eq. - rewrite fst_split_map_eq. - rewrite firstn_map. - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply allNatsLt_length. - eapply list_pred_eq_gen. - - symmetry. - eapply (@firstn_allNatsLt _ (S maxMatches)). - omega. - eapply list_pred_I. - rewrite (@firstn_allNatsLt _ (S maxMatches)). - symmetry. - eapply allNatsLt_length. - omega. - intuition. - simpl in *; subst. - destruct (F b0 b2). - destruct p. - trivial. - - intuition. - - eapply in_split_l_if in H5. - destruct H5. - eapply in_flatten in H5. - destruct H5. - intuition. - eapply in_map_iff in H6. - destruct H6. - intuition. - destruct x2. - destruct p. - subst. - unfold ESPADA_TSet_Once_Games.numberedMap in *. - eapply in_map_iff in H7. - destruct H7. - intuition. - destruct x1. - - eapply in_fst_split_if. - eapply in_flatten. - econstructor. - split. - eapply in_map_iff. - econstructor. - split. - reflexivity. - eauto. - simpl. - eapply in_map_iff. - econstructor. - split; eauto. - simpl. - remember (F b0 n) as z. - destruct z. - destruct p. - pairInv. - - Theorem labelCollision_9_coll_eq : - forall ls a1 a2 i1 i2, - (NoDup (getAllLabels _ maxMatches F ls) -> - In a1 ls -> - In a2 ls -> - i1 <= maxMatches -> - i2 <= maxMatches -> - fst (F a1 i1) = fst (F a2 i2) -> - (a1, i1) = (a2, i2))%nat. - - induction ls; intuition; simpl in *. - intuition. - - intuition; subst. - subst. - - eapply NoDup_app in H0. - intuition; subst. - - assert (i1 = i2). - eapply NoDup_map. - eapply H1. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - simpl. - trivial. - subst. - trivial. - - - eapply NoDup_app in H0. - intuition; subst. - - exfalso. - eapply H7. - eapply in_map_iff. - exists i1. - split. - reflexivity. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - Focus 2. - eapply H5. - eapply In_getAllLabels; - eauto. - - eapply NoDup_app in H0. - intuition; subst. - - exfalso. - eapply H7. - eapply in_map_iff. - exists i2. - split. - reflexivity. - eapply (@allNatsLt_lt_if (S maxMatches)). - omega. - Focus 2. - symmetry. - eapply H5. - eapply In_getAllLabels; - eauto. - - eapply NoDup_app in H0. - intuition; subst. - - Qed. - - assert ((b0, n) = (a, i)). - eapply labelCollision_9_coll_eq; eauto. - eapply hasDups_false_NoDup. - eauto. - eapply In_zip in H8. - intuition. - eapply In_zip in H3. - intuition. - eapply In_zip in H7. - intuition. - - eapply In_zip in H8. - intuition. - eapply In_zip in H7. - intuition. - eapply in_map_iff in H10. - destruct H10. - intuition. - subst. - eapply allNatsLt_lt in H3. - eapply le_trans. - eapply Nat.lt_le_incl. - eapply H3. - - eapply arrayLookupList_pred. - eapply maxMatches_correct. - eauto. - simpl. - intuition. - rewrite <- Heqz. - trivial. - pairInv. - eauto. - Qed. - - Theorem ESPADA_TSetCorr_G12_G13_equiv : - Pr[ESPADA_TSetCorr_G12 maxMatches F A1 ] == Pr[ESPADA_TSetCorr_G13 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G12, ESPADA_TSetCorr_G13. - - comp_skip. - comp_simp. - comp_skip. - eapply evalDist_ret_eq. - - case_eq (labelCollision_9 lambda maxMatches F x); intuition. - simpl. - f_equal. - f_equal. - - rewrite (@map_snd_eq _ (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t))). - symmetry. - rewrite (@map_fst_eq _ x). - symmetry. - eapply map_ext_in. - intuition. - - destruct a. - simpl. - - - Theorem TSetRetrieve_tLoop_Corr_12_13_equiv : - forall i (t1 : TSetCorr_12 lambda) (t2 : TSetCorr_13 lambda) m x y, - (forall a b c, In (a, c) m -> In (b, c) m -> a = b) -> - (forall a b c, In (a, b) m -> In (a, c) m -> b = c) -> - In (y, x) m -> - NoDup (fst (split t2)) -> - list_pred (fun a b => - let '(t1, i1, (q1, l1, s1)) := a in - let '(q2, i2, (l2, s2)) := b in - q1 = q2 /\ i1 = i2 /\ l1 = l2 /\ s1 = s2 /\ - In (q1, t1) m) t1 t2 - -> - ESPADA_TSetRetrieve_tLoop_Corr_12 t1 x i = - ESPADA_TSetRetrieve_tLoop_Corr_13 t2 y i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_12, ESPADA_TSetRetrieve_tLoop_Corr_13. - case_eq ( arrayLookup (pair_EqDec (Bvector_EqDec lambda) nat_EqDec) t1 (x, i)); intuition. - eapply arrayLookup_Some_impl_In in H. - specialize (list_pred_In_exists H4 _ H); intuition. - destruct H6. - intuition. - destruct p. - destruct p. - destruct x0. - destruct p. - destruct p0. - intuition. - subst. - assert (y = b1). - eapply H0; eauto. - subst. - erewrite arrayLookup_In_NoDup; eauto. - comp_simp. - trivial. - - comp_simp. - case_eq ( @arrayLookup (prod Blist nat) - (prod nat (Bvector (posnatToNat lambda))) - (@pair_EqDec Blist nat (@list_EqDec bool bool_EqDec) nat_EqDec) t2 - (@pair Blist nat y i)); intuition. - destruct p. - eapply list_pred_symm in H4. - eapply arrayLookup_Some_impl_In in H6. - specialize (list_pred_In_exists H4 _ H6); intuition. - destruct H7. - intuition. - destruct x0. - destruct p. - destruct p0. - destruct p. - intuition. - subst. - assert (x = b0). - eapply H1; eauto. - subst. - eapply arrayLookup_None_not_In in H. - intuition. - eapply in_fst_split_if. - eauto. - Qed. - - Theorem TSetRetrieve_h_Corr_12_13_equiv : - forall fuel i t1 t2 m x y, - (forall a b c, In (a, c) m -> In (b, c) m -> a = b) -> - (forall a b c, In (a, b) m -> In (a, c) m -> b = c) -> - In (y, x) m -> - NoDup (fst (split t2)) -> - list_pred (fun a b => - let '(t1, i1, (q1, l1, s1)) := a in - let '(q2, i2, (l2, s2)) := b in - q1 = q2 /\ i1 = i2 /\ l1 = l2 /\ s1 = s2 /\ - In (q1, t1) m) t1 t2 - -> - (@ESPADA_TSetRetrieve_h_Corr_12 lambda) t1 x i fuel = - ESPADA_TSetRetrieve_h_Corr_13 t2 y i fuel. - - induction fuel; intuition; simpl in *. - rewrite (@TSetRetrieve_tLoop_Corr_12_13_equiv _ _ t2 m _ y). - destruct (ESPADA_TSetRetrieve_tLoop_Corr_13 t2 y i). - destruct p. - f_equal. - destruct b0. - eauto. - trivial. - trivial. - - eauto. - eauto. - trivial. - trivial. - trivial. - - Qed. - - Theorem TSetRetrieve_Corr_12_13_equiv : - forall t1 t2 m x y, - (forall a b c, In (a, c) m -> In (b, c) m -> a = b) -> - (forall a b c, In (a, b) m -> In (a, c) m -> b = c) -> - In (y, x) m -> - NoDup (fst (split t2)) -> - list_pred (fun a b => - let '(t1, i1, (q1, l1, s1)) := a in - let '(q2, i2, (l2, s2)) := b in - q1 = q2 /\ i1 = i2 /\ l1 = l2 /\ s1 = s2 /\ - In (q1, t1) m) t1 t2 - -> - (@ESPADA_TSetRetrieve_Corr_12 lambda) maxMatches t1 x = - ESPADA_TSetRetrieve_Corr_13 maxMatches t2 y. - - intuition. - - eapply TSetRetrieve_h_Corr_12_13_equiv; eauto. - - Qed. - - Show. - eapply TSetRetrieve_Corr_12_13_equiv; eauto. - - intuition. - - - - eapply In_combine_NoDup_eq_l; eauto. - - Theorem NoDup_labels_NoDup : - forall ls, - NoDup (getAllLabels _ maxMatches F ls) -> - NoDup ls. - - induction ls; intuition; simpl in *. - econstructor. - - eapply NoDup_app in H. - intuition. - - econstructor. - intuition. - eapply (@H3 (fst (F a maxMatches)) (fst (F a maxMatches))). - eapply in_map_iff. - exists maxMatches. - intuition. - - eapply In_getAllLabels; intuition. - trivial. - trivial. - Qed. - - eapply NoDup_labels_NoDup . - eapply hasDups_false_NoDup. - trivial. - - intuition. - eapply In_combine_NoDup_eq_r; eauto. - eapply combineKeywords_NoDup. - eapply removeDups_NoDup. - eapply toW_NoDup. - - rewrite fst_split_flatten_eq. - - rewrite map_map. - - Require Import NoDup_gen. - - - eapply flatten_NoDup_gen. - - eapply (@map_NoDup_gen _ _ (fun a b => (snd a) <> nil /\ length (snd a) = length (snd b) /\ fst (fst a) = fst (fst b))); intros. - intuition. - destruct a1; destruct a2; simpl in *. - subst. - destruct l0; simpl in *; try omega. - intuition. - intuition. - destruct a1; destruct a2; destruct a3; simpl in *. - destruct p; destruct p0; destruct p1; simpl in *. - subst. - trivial. - - destruct a1; destruct a2; simpl in *. - destruct p; destruct p0; simpl in *. - intuition; subst. - unfold ESPADA_TSet_Once_Games.numberedMap in *. - rewrite fst_split_map_eq in H5. - eapply map_eq_nil in H5. - eapply H4. - - eapply zip_eq_nil_l in H5. - destruct l0; simpl in *; trivial. - eapply app_eq_nil in H5. - simpl in *; intuition; discriminate. - eapply allNatsLt_length. - - unfold ESPADA_TSet_Once_Games.numberedMap. - repeat rewrite fst_split_map_eq. - assert ( - map - (fun a : nat * Bvector lambda => - fst ([i, a0]<-2 a; [_, _, _]<-3 F b2 i; (b3, i, (length l0, a0)))) - (zip (allNatsLt (length l0)) l0) - = - map - (fun a => (b3, a)) - (allNatsLt (length l0)) - ). - - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply allNatsLt_length. - eapply list_pred_eq. - eapply list_pred_I. - symmetry. - eapply allNatsLt_length. - intuition. - subst. - comp_simp. - trivial. - rewrite H5. - clear H5. - - assert ( - map - (fun a : nat * Bvector lambda => - fst ([i, a0]<-2 a; [_, _, _]<-3 F b4 i; (b3, i, (length l1, a0)))) - (zip (allNatsLt (length l1)) l1) - = - map - (fun a => - (b3, a)) - (allNatsLt (length l1)) - ). - - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply allNatsLt_length. - eapply list_pred_eq. - eapply list_pred_I. - symmetry. - eapply allNatsLt_length. - intuition. - subst. - comp_simp. - trivial. - rewrite H5. - clear H5. - rewrite H3. - trivial. - - intuition. - subst. - intuition. - - intuition. - subst. - trivial. - - destruct a. - destruct p. - simpl in *. - intuition. - eapply H4. - unfold numberedMap. - subst. - simpl. - trivial. - - destruct a1. - destruct p. - destruct a2. - destruct p. - simpl in *. - unfold ESPADA_TSet_Once_Games.numberedMap in *. - repeat rewrite fst_split_map_eq in *. - - assert ( - map - (fun a : nat * Bvector lambda => - fst ([i, a0]<-2 a; [_, _, _]<-3 F b2 i; (b1, i, (length l0, a0)))) - (zip (allNatsLt (length l0)) l0) - = - map - (fun a => - (b1, a)) - (allNatsLt (length l0)) - ). - - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply allNatsLt_length. - eapply list_pred_eq. - eapply list_pred_I. - symmetry. - eapply allNatsLt_length. - intuition. - subst. - comp_simp. - trivial. - rewrite H4 in H3. - clear H4. - - assert ( - map - (fun a : nat * Bvector lambda => - fst ([i, a0]<-2 a; [_, _, _]<-3 F b4 i; (b3, i, (length l1, a0)))) - (zip (allNatsLt (length l1)) l1) - = - map - (fun a => - (b3, a)) - (allNatsLt (length l1)) - ). - - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply allNatsLt_length. - eapply list_pred_eq. - eapply list_pred_I. - symmetry. - eapply allNatsLt_length. - intuition. - subst. - comp_simp. - trivial. - rewrite H4 in H3. - clear H4. - - intuition. - subst. - eapply H4. - trivial. - - rewrite <- (@allNatsLt_length (length l0)). - rewrite <- (@map_length _ _ (fun a : nat => (b1, a))). - rewrite H5. - rewrite map_length. - eapply allNatsLt_length. - - eapply map_pair_fst_eq . - eapply H5. - destruct l0; simpl in *; intuition. - eapply app_eq_nil in H3. - simpl in *; intuition; discriminate. - - eapply (@NoDup_gen_weaken _ (fun a b => fst (fst a) = fst (fst b))). - - eapply (@NoDup_gen_zip_fst _ _ (fun a b => fst a = fst b)). - eapply NoDup_gen_zip_fst. - eapply NoDup_gen_eq . - eapply combineKeywords_NoDup. - eapply removeDups_NoDup. - eapply toW_NoDup. - - intuition. - intuition. - eapply in_map_iff in H3. - destruct H3. - intuition. - subst. - unfold ESPADA_TSet_Once_Games.numberedMap. - destruct x1. - destruct p. - rewrite fst_split_map_eq. - eapply NoDup_gen_eq. - eapply (@map_NoDup_gen _ _ (fun a b => fst a = fst b)); - intuition; simpl in *; subst; trivial. - comp_simp. - trivial. - destruct a2. - destruct (F b2 a). - destruct p. - destruct (F b2 n). - destruct p. - simpl in *. - pairInv. - trivial. - - eapply NoDup_gen_zip_fst. - eapply NoDup_gen_eq . - eapply allNatsLt_NoDup. - - intuition. - eapply app_NoDup. - eapply in_map_iff in H3. - destruct H3. - intuition. - subst. - comp_simp. - unfold ESPADA_TSet_Once_Games.numberedMap. - rewrite fst_split_map_eq. - eapply NoDup_gen_eq. - eapply (@map_NoDup_gen _ _ (fun a b => fst a = fst b)); - intuition; simpl in *; subst; trivial. - comp_simp. - trivial. - destruct (F b2 a). - destruct p. - destruct a2. - destruct (F b2 n0). - destruct p. - simpl in *. - pairInv. - trivial. - eapply NoDup_gen_zip_fst. - eapply NoDup_gen_eq . - eapply allNatsLt_NoDup. - - repeat intac; subst. - unfold ESPADA_TSet_Once_Games.numberedMap. - destruct x0. - destruct p. - rewrite fst_split_map_eq. - eapply NoDup_gen_eq. - eapply (@map_NoDup_gen _ _ (fun a b => fst a = fst b)); - intuition; simpl in *; subst; trivial. - comp_simp. - trivial. - destruct a2. - destruct (F b2 a). - destruct p. - destruct (F b2 n). - destruct p. - simpl in *. - pairInv. - trivial. - - eapply NoDup_gen_zip_fst. - eapply NoDup_gen_eq . - eapply allNatsLt_NoDup. - - intuition. - eapply H6. - clear H6. - - repeat intac; subst. - - destruct x3. - destruct p. - destruct x0. - destruct p. - unfold ESPADA_TSet_Once_Games.numberedMap in *. - repeat rewrite fst_split_map_eq in *. - eapply in_map_iff in H8. - eapply in_map_iff in H7. - destruct H8. - destruct H7. - intuition. - subst. - destruct x1. - destruct x0. - destruct (F b2 n). - destruct (F b4 n0). - destruct p. - destruct p0. - simpl in *. - pairInv. - - Theorem In_zip_3 : - forall (A B C : Set)(lsa : list A)(lsb : list B)(lsc : list C) a b c, - In (a, b, c) (zip (zip lsa lsb) lsc) -> - In (a, c) (zip lsa lsc). - - induction lsa; intuition; simpl in *. - destruct lsb; - simpl in *. - intuition. - - destruct lsc; simpl in *. - intuition. - - intuition. - pairInv. - intuition. - right. - eapply IHlsa. - eauto. - - Qed. - - eapply In_zip_3 in H6. - eapply In_zip_3 in H9. - eapply In_zip_strong in H6. - eapply In_zip_strong in H9. - intuition. - subst. - eapply map_ext_in. - intuition. - comp_simp. - trivial. - - intuition. - repeat intac; subst. - eapply H6. - clear H6. - comp_simp. - eapply In_zip_3 in H10. - eapply In_zip_3 in H9. - eapply In_zip_strong in H10. - eapply In_zip_strong in H9. - intuition. - subst. - unfold ESPADA_TSet_Once_Games.numberedMap in *. - repeat rewrite fst_split_map_eq in *. - repeat intac; subst. - destruct x0. - destruct (F b2 n). - destruct p. - destruct x1. - destruct (F b4 n1). - destruct p. - simpl in *. - pairInv. - eapply map_ext_in. - intuition. - comp_simp. - trivial. - - eapply list_pred_flatten_both. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - - eapply list_pred_impl. - - eapply list_pred_eq_in. - intuition. - subst. - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq_in. - intuition. - subst. - destruct (F b2 a1). - destruct p. - intuition. - - eapply In_zip in H7. - intuition. - - rewrite <- zip_combine_eq. - eauto. - - symmetry. - eapply compMap_length. - eauto. - - eapply compMap_length. - eauto. - Qed. - - Definition ESPADA_TSetCorr_G13_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map - (fun p => [q, ls] <-2 p; numberedMap - (fun i len s_i => - ((q, i), (len, s_i))) ls) (zip allQs ts); - tSet <- (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allQs)ESPADA_TSetRetrieve_Corr_13 maxMatches tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 _ maxMatches F allTags) || bad). - - Definition TSetCorr_13_2_Record := (nat * (nat * (Bvector lambda)))%type. - - Definition TSetCorr_13_2 := list (Blist * list TSetCorr_13_2_Record). - - - Definition ESPADA_TSetRetrieve_tLoop_Corr_13_2 (tSet : list TSetCorr_13_2_Record) i := - t <- arrayLookup _ tSet i; - match t with - | None => None - | Some (len, s_i) => - Some (s_i, if (eq_nat_dec (S i) len) then false else true) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_13_2 (tSet : list TSetCorr_13_2_Record) i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_13_2 tSet i) with - | Some (v, b) => - v :: (if b then (ESPADA_TSetRetrieve_h_Corr_13_2 tSet (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_13_2 tSet (q : Blist) := - ls <- arrayLookupList _ tSet q; - ESPADA_TSetRetrieve_h_Corr_13_2 ls O maxMatches. - - Definition ESPADA_TSetCorr_G13_2 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map - (fun p => [q, ls] <-2 p; numberedMap - (fun i len s_i => - (i, (len, s_i))) ls) (zip allQs ts); - tSet <- combine allQs ts_recsList ; - bad <- negb - (eqb (foreach (x in allQs)ESPADA_TSetRetrieve_Corr_13_2 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 _ maxMatches F allTags) || bad). - - Theorem ESPADA_TSetCorr_G13_1_equiv : - Pr[ESPADA_TSetCorr_G13 maxMatches F A1] == Pr[ESPADA_TSetCorr_G13_1]. - - unfold ESPADA_TSetCorr_G13, ESPADA_TSetCorr_G13_1. - - comp_skip. - comp_simp. - comp_skip. - - eapply evalDist_ret_eq. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - f_equal. - f_equal. - eapply map_eq. - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_map_r. - eapply list_pred_zip_l. - eapply list_pred_I. - symmetry. - eapply compMap_length. - eauto. - eapply list_pred_eq. - eapply list_pred_I. - eapply compMap_length. - eauto. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_zip_r. - eapply list_pred_I. - symmetry. - eapply compMap_length. - eauto. - eapply list_pred_eq. - eapply list_pred_I. - eapply compMap_length. - eauto. - eapply list_pred_map_l. - eapply list_pred_zip_r. - eapply list_pred_I. - symmetry. - eapply compMap_length. - eauto. - eapply list_pred_eq. - eapply list_pred_I. - eapply compMap_length. - eauto. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_eq. - intuition. - repeat exdest. - simpl in *. - subst. - intuition. - subst. - destruct b1. - simpl in *. - subst. - unfold numberedMap. - eapply map_ext_in. - intuition. - comp_simp. - trivial. - Qed. - - Theorem ESPADA_TSetCorr_G13_1_2_equiv : - Pr[ESPADA_TSetCorr_G13_1] == Pr[ESPADA_TSetCorr_G13_2]. - - unfold ESPADA_TSetCorr_G13_1, ESPADA_TSetCorr_G13_2. - - comp_skip; comp_simp. - comp_skip. - eapply evalDist_ret_eq. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - - Theorem TSetRetrieve_tLoop_Corr_13_2_equiv : - forall i ls1 ls2 q, - (arrayLookup _ (flatten ls1) (q, i) = - arrayLookup _ (arrayLookupList _ ls2 q) i) -> - ESPADA_TSetRetrieve_tLoop_Corr_13 (flatten ls1) q i = - ESPADA_TSetRetrieve_tLoop_Corr_13_2 - (arrayLookupList (list_EqDec bool_EqDec) ls2 q) i. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_13, ESPADA_TSetRetrieve_tLoop_Corr_13_2. - rewrite H. - comp_simp. - unfold Blist in *. - destruct ( arrayLookup nat_EqDec (arrayLookupList (list_EqDec bool_EqDec) ls2 q) i); intuition. - Qed. - - - Theorem TSetRetrieve_h_Corr_13_2_equiv : - forall fuel i ls1 ls2 q, - (forall i, (arrayLookup _ (flatten ls1) (q, i) = - arrayLookup _ (arrayLookupList _ ls2 q) i)) -> - ESPADA_TSetRetrieve_h_Corr_13 (flatten ls1) q i fuel = - ESPADA_TSetRetrieve_h_Corr_13_2 - (arrayLookupList (list_EqDec bool_EqDec) ls2 q) i fuel. - - induction fuel; intuition; simpl in *. - rewrite (@TSetRetrieve_tLoop_Corr_13_2_equiv _ _ ls2) . - destruct (ESPADA_TSetRetrieve_tLoop_Corr_13_2 - (arrayLookupList (list_EqDec bool_EqDec) ls2 q) i). - comp_simp. - f_equal. - destruct b0; trivial. - eapply IHfuel; intuition. - trivial. - eauto. - - Qed. - - Theorem TSetRetrieve_Corr_13_2_equiv : - forall ls1 ls2 q, - (forall i, (arrayLookup _ (flatten ls1) (q, i) = - arrayLookup _ (arrayLookupList _ ls2 q) i)) -> - ESPADA_TSetRetrieve_Corr_13 maxMatches (flatten ls1) q = - ESPADA_TSetRetrieve_Corr_13_2 ls2 q. - - intuition. - unfold ESPADA_TSetRetrieve_Corr_13, ESPADA_TSetRetrieve_Corr_13_2. - unfold setLet. - apply TSetRetrieve_h_Corr_13_2_equiv. - eauto. - - Qed. - - eapply TSetRetrieve_Corr_13_2_equiv. - intuition. - - - erewrite (@arrayLookup_flatten_eq _ _ _ _ _ _ (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t))). - - eapply arrayLookup_pair_snd . - - eapply arrayLookup_pred_2 . - repeat rewrite <- zip_combine_eq. - eapply list_pred_impl. - - Ltac predtac := - match goal with - | [ |- _ ] => eapply list_pred_eq - | [ |- list_pred _ _ (map _ _ ) ] => eapply list_pred_map_r - | [ |- list_pred _ (map _ _ ) _ ] => eapply list_pred_map_l - | [ |- list_pred _ _ (zip _ _ ) ] => eapply list_pred_zip_r - | [ |- list_pred _ (zip _ _ ) _ ] => eapply list_pred_zip_l - end. - - repeat predtac. - - intuition. - repeat (exdest; intuition); simpl in *; subst. - subst. - trivial. - subst. - repeat (exdest; intuition); simpl in *; subst. - destruct x0. - destruct b0. - simpl in *. - subst. - destruct x6. - unfold numberedMap. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - simpl in *. - subst. - eapply list_pred_impl. - eapply list_pred_eq. - intuition. - subst. - simpl. - trivial. - subst. - simpl. - trivial. - repeat (exdest; intuition); simpl in *; subst. - trivial. - repeat (exdest; intuition); simpl in *; subst. - destruct x0. - destruct b0. - simpl in *; subst. - destruct x6. - simpl in *. - subst. - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - - intuition. - subst. - simpl. - trivial. - subst. - simpl. - trivial. - econstructor. - - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_eq. - eapply list_pred_map_l. - eapply list_pred_eq. - - intuition. - intuition. - repeat (exdest; intuition). - subst. - destruct b. - destruct x0. - destruct p. - simpl in *. - unfold numberedMap in *. - eapply in_map_iff in H3. - exdest. - intuition. - destruct x0. - pairInv. - trivial. - - eapply combineKeywords_NoDup. - eapply removeDups_NoDup. - eapply toW_NoDup. - - Qed. - - Definition TSetCorr_13_3_Record := (nat * ( (Bvector lambda)))%type. - - Definition TSetCorr_13_3 := list (Blist * list TSetCorr_13_3_Record). - - - Definition ESPADA_TSetRetrieve_tLoop_Corr_13_3 (tSet : list TSetCorr_13_3_Record) i := - arrayLookup _ tSet i. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_13_3 (tSet : list TSetCorr_13_3_Record) i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_13_3 tSet i) with - | Some v => - v :: (ESPADA_TSetRetrieve_h_Corr_13_3 tSet (S i) fuel') - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_13_3 tSet (q : Blist) := - ls <- arrayLookupList _ tSet q; - ESPADA_TSetRetrieve_h_Corr_13_3 ls O maxMatches. - - Definition ESPADA_TSetCorr_G13_3 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map - (fun p => [q, ls] <-2 p; numberedMap - (fun i len s_i => - (i, s_i)) ls) (zip allQs ts); - tSet <- combine allQs ts_recsList ; - bad <- negb - (eqb (foreach (x in allQs)ESPADA_TSetRetrieve_Corr_13_3 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 _ maxMatches F allTags) || bad). - - - Theorem ESPADA_TSetCorr_G13_2_3_equiv : - Pr[ESPADA_TSetCorr_G13_2] == Pr[ESPADA_TSetCorr_G13_3]. - - unfold ESPADA_TSetCorr_G13_2, ESPADA_TSetCorr_G13_3. - comp_skip; comp_simp. - comp_skip. - eapply evalDist_ret_eq. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - - Theorem TSetRetrieve_tLoop_13_2_3_equiv : - forall i t1 t2, - (arrayLookup _ t2 i = - match (arrayLookup _ t1 i) with - | Some (_, b) => Some b - | None => None - end) -> - ESPADA_TSetRetrieve_tLoop_Corr_13_3 t2 i = - match (ESPADA_TSetRetrieve_tLoop_Corr_13_2 t1 i) with - | Some (x, b) => Some x - | None => None - end. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_13_3, ESPADA_TSetRetrieve_tLoop_Corr_13_2. - unfold setLet. - rewrite H. - case_eq (arrayLookup nat_EqDec t1 i); intuition. - destruct p. - trivial. - - Qed. - - - Theorem TSetRetrieve_h_13_2_3_equiv : - forall fuel i t1 t2, - list_pred (fun x y => fst x = fst y /\ snd (snd x) = (snd y) /\ - (S (fst x) = fst (snd x) -> arrayLookup _ t2 (S (fst x)) = None)) t1 t2 -> - NoDup (fst (split t1)) -> - (forall n, In n (fst (split (snd (split t1)))) -> n = length t1) -> - ESPADA_TSetRetrieve_h_Corr_13_2 t1 i fuel = - ESPADA_TSetRetrieve_h_Corr_13_3 t2 i fuel. - - induction fuel; intuition; simpl in *. - unfold ESPADA_TSetRetrieve_tLoop_Corr_13_2, ESPADA_TSetRetrieve_tLoop_Corr_13_3. - unfold setLet. - - case_eq (arrayLookup nat_EqDec t1 i); intuition. - destruct p. - eapply arrayLookup_Some_impl_In in H2. - specialize (list_pred_In_exists H _ H2); intuition. - destruct H3. - intuition. - destruct (eq_nat_dec (S i) n). - simpl in *. - subst. - destruct x. - simpl in *. - erewrite arrayLookup_In_NoDup; eauto. - f_equal. - destruct fuel; simpl in *; trivial. - unfold ESPADA_TSetRetrieve_tLoop_Corr_13_3. - rewrite H8. - trivial. - trivial. - - rewrite (@list_pred_fst_split_eq _ _ _ _ t1). - trivial. - eapply list_pred_symm. - eapply list_pred_impl. - eauto. - intuition. - - simpl in *. - subst. - destruct x. - simpl in *. - erewrite arrayLookup_In_NoDup; eauto. - f_equal. - eauto. - rewrite (@list_pred_fst_split_eq _ _ _ _ t1). - trivial. - eapply list_pred_symm. - eapply list_pred_impl. - eauto. - intuition. - - rewrite notInArrayLookupNone. - trivial. - intuition. - - rewrite unzip_eq_split in H3. - rewrite (@list_pred_fst_split_eq _ _ _ _ t1) in H3. - eapply arrayLookup_None_not_In. - eauto. - trivial. - eapply list_pred_symm. - eapply list_pred_impl. - eauto. - intuition. - Qed. - - Theorem TSetRetrieve_13_2_3_equiv : - forall t1 t2 a, - (list_pred - (fun (x : nat * (nat * Bvector lambda)) (y : nat * Bvector lambda) => - fst x = fst y /\ - snd (snd x) = snd y /\ - (S (fst x) = fst (snd x) -> - arrayLookup nat_EqDec (arrayLookupList (list_EqDec bool_EqDec) t2 a) - (S (fst x)) = None)) (arrayLookupList (list_EqDec bool_EqDec) t1 a) - (arrayLookupList (list_EqDec bool_EqDec) t2 a)) -> - - NoDup (fst (split (arrayLookupList (list_EqDec bool_EqDec) t1 a))) -> - - (forall n : nat, - In n - (fst - (split (snd (split (arrayLookupList (list_EqDec bool_EqDec) t1 a))))) -> - n = length (arrayLookupList (list_EqDec bool_EqDec) t1 a)) -> - - ESPADA_TSetRetrieve_Corr_13_2 t1 a = - ESPADA_TSetRetrieve_Corr_13_3 t2 a. - - intuition. - unfold ESPADA_TSetRetrieve_Corr_13_2, ESPADA_TSetRetrieve_Corr_13_3. - unfold setLet. - eapply TSetRetrieve_h_13_2_3_equiv ; - eauto. - Qed. - - eapply TSetRetrieve_13_2_3_equiv. - intuition. - - eapply arrayLookup_pred_2. - repeat rewrite <- zip_combine_eq. - eapply list_pred_impl. - repeat predtac. - intuition. - repeat (exdest; intuition); simpl in *; subst. - trivial. - repeat (exdest; intuition); simpl in *; subst. - destruct b0; simpl in *; subst. - destruct x0; destruct x6. - unfold numberedMap. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - simpl in *; subst. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - trivial. - simpl. - trivial. - - simpl in *. - subst. - eapply notInArrayLookupNone. - intuition. - - rewrite unzip_eq_split in H2. - eapply in_split_l_if in H2. - destruct H2. - - destruct x13; simpl in *; subst. - destruct x3; simpl in *; subst. - destruct x10; simpl in *; subst. - - eapply In_arrayLookupList in H2. - destruct H2. - intuition. - eapply list_pred_zip_in in H5. - Focus 2. - eapply list_pred_map_r. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_eq. - eapply list_pred_map_l. - eapply list_pred_eq. - simpl in *. - repeat (exdest; intuition); simpl in *; subst. - destruct x2. - eapply in_map_iff in H6. - destruct H6. - intuition. - destruct x1. - pairInv. - simpl in *. - subst. - eapply In_zip in H6. - intuition. - - eapply allNatsLt_lt in H2. - omega. - - repeat (exdest; intuition). simpl in *; subst. - trivial. - - repeat (exdest; intuition). simpl in *; subst. - destruct b0; simpl in *. - subst. - destruct x0. - destruct x6. - simpl in *; subst. - unfold numberedMap. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - trivial. - trivial. - - simpl in *; subst. - destruct x3, x13, x10; simpl in *; subst. - eapply notInArrayLookupNone. - intuition. - rewrite unzip_eq_split in H3. - eapply in_split_l_if in H3. - destruct H3. - eapply In_arrayLookupList in H2. - destruct H2. - intuition. - eapply list_pred_zip_in in H3. - Focus 2. - eapply list_pred_map_r. - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_eq. - eapply list_pred_eq. - eapply list_pred_map_l. - eapply list_pred_eq. - simpl in *. - repeat (exdest; intuition); simpl in *; subst. - destruct x2. - eapply in_map_iff in H6. - destruct H6. - intuition. - destruct x1. - pairInv. - simpl in *. - subst. - eapply In_zip in H6. - intuition. - - eapply allNatsLt_lt in H2. - omega. - - econstructor. - eapply arrayLookupList_pred. - intuition. - destruct x0. - simpl. - eapply in_combine_r in H2. - repeat intac. - destruct x0. - subst. - unfold numberedMap. - rewrite fst_split_map_eq. - eapply NoDup_gen_eq. - eapply (@map_NoDup_gen _ _ (fun a b => fst a = fst b)); intuition; simpl in *; subst. - destruct a2; simpl in *; trivial. - destruct a2; simpl in *; trivial. - eapply NoDup_gen_zip_fst. - eapply NoDup_gen_eq. - eapply allNatsLt_NoDup. - simpl. - econstructor. - - intuition. - eapply in_split_l_if in H2. - destruct H2. - - eapply in_split_r_if in H2. - destruct H2. - eapply In_arrayLookupList in H2. - destruct H2. - intuition. - rewrite <- zip_combine_eq in H3. - eapply list_pred_zip_in in H3. - Focus 2. - repeat predtac. - simpl in *. - repeat (exdest; intuition). - simpl in *; subst. - destruct x3. - simpl in *; subst. - unfold numberedMap in *. - repeat intac. - destruct x2. - pairInv. - - symmetry. - eapply arrayLookupList_pred'. - intuition. - rewrite <- zip_combine_eq in H2. - eapply list_pred_zip_in in H2. - Focus 2. - repeat predtac. - simpl in *. - repeat (exdest; intuition). - subst. - destruct x3; simpl in *; subst. - rewrite map_length. - rewrite zip_length. - eapply allNatsLt_length. - eapply allNatsLt_length. - intuition. - eapply arrayLookup_None_not_In in H2; - intuition. - rewrite combine_split. - simpl. - trivial. - rewrite map_length. - rewrite zip_length. - trivial. - rewrite map_length. - trivial. - Qed. - - Theorem ESPADA_TSetCorr_G13_3_G14_equiv : - Pr[ESPADA_TSetCorr_G13_3] == Pr[ESPADA_TSetCorr_G14 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G13_3, ESPADA_TSetCorr_G14. - comp_skip; comp_simp. - comp_skip. - eapply evalDist_ret_eq. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - - unfold ESPADA_TSetRetrieve_Corr_13_3, ESPADA_TSetRetrieve_Corr_14. - unfold setLet. - - Theorem skipn_length_nil : - forall (A : Type)(ls : list A) n, - n >= length ls -> - skipn n ls = nil. - - induction ls; destruct n; simpl in *; intuition. - omega. - - Qed. - - Theorem ESPADA_TSetRetrieve_tLoop_Corr_13_3_eq : - forall ls i, - fst (split ls) = allNatsLt (length ls) -> - ESPADA_TSetRetrieve_tLoop_Corr_13_3 ls i = - match (nth_option ls i) with - | None => None - | Some (_, v) => Some v - end. - - intuition. - unfold ESPADA_TSetRetrieve_tLoop_Corr_13_3. - - eapply arrayLookup_allNats_eq . - trivial. - - Qed. - - Theorem ESPADA_TSetRetrieve_h_Corr_13_3_eq : - forall fuel i ls, - fst (split ls) = allNatsLt (length ls) -> - fuel + i >= length ls-> - ESPADA_TSetRetrieve_h_Corr_13_3 ls i fuel = - skipn i (snd (split ls)). - - induction fuel; intuition; simpl in *. - symmetry. - eapply skipn_length_nil. - rewrite split_length_r. - trivial. - - rewrite ESPADA_TSetRetrieve_tLoop_Corr_13_3_eq. - case_eq ( @nth_option (prod nat (Bvector (posnatToNat lambda))) ls i); intuition. - destruct p. - rewrite IHfuel. - - symmetry. - eapply skipn_S_eq . - - eapply nth_option_snd_split. - eauto. - trivial. - omega. - - eapply nth_option_None_ge in H1. - rewrite skipn_length_nil. - trivial. - rewrite split_length_r. - trivial. - - trivial. - - Qed. - - rewrite ESPADA_TSetRetrieve_h_Corr_13_3_eq. - simpl. - eapply arrayLookupList_pred'. - intuition. - repeat rewrite <- zip_combine_eq in H2. - eapply list_pred_zip_in in H2. - Focus 2. - repeat predtac. - simpl in *. - repeat (exdest; intuition). - subst. - destruct x1. - simpl in *; subst. - symmetry. - eapply arrayLookupList_pred'. - intuition. - repeat rewrite <- zip_combine_eq in H2. - eapply list_pred_zip_in in H2. - Focus 2. - eapply list_pred_map_r. - eapply list_pred_eq. - simpl in *. - exdest. - intuition. - subst. - unfold numberedMap. - - rewrite snd_split_map_eq. - - assert ( - (@map (prod nat (Bvector (posnatToNat lambda))) - (Bvector (posnatToNat lambda)) - (fun p : prod nat (Bvector (posnatToNat lambda)) => - @snd nat (Bvector (posnatToNat lambda)) - match p return (prod nat (Bvector (posnatToNat lambda))) with - | pair i a => @pair nat (Bvector (posnatToNat lambda)) i a - end) - (@zip nat (Bvector (posnatToNat lambda)) - (allNatsLt - (@length (Bvector (posnatToNat lambda)) - (@arrayLookupList Blist (Bvector (posnatToNat lambda)) - (@list_EqDec bool bool_EqDec) t x1))) - (@arrayLookupList Blist (Bvector (posnatToNat lambda)) - (@list_EqDec bool bool_EqDec) t x1))) - = - map (@snd _ _ ) - (zip (allNatsLt (length (arrayLookupList (list_EqDec bool_EqDec) t x1))) - (arrayLookupList (list_EqDec bool_EqDec) t x1)) - ). - eapply map_ext. - intuition. - rewrite H2. - clear H2. - rewrite map_snd_zip. - trivial. - eapply allNatsLt_length. - - intuition. - eapply arrayLookup_None_not_In in H2. - intuition. - rewrite combine_split. - simpl. - trivial. - - rewrite map_length. - trivial. - - intuition. - eapply arrayLookup_None_not_In in H2. - intuition. - rewrite combine_split. - simpl. - trivial. - - rewrite map_length. - rewrite zip_length. - trivial. - rewrite map_length. - trivial. - - eapply arrayLookupList_pred'. - intuition. - repeat rewrite <- zip_combine_eq in H2. - repeat intac. - subst. - destruct x1. - unfold numberedMap. - assert ( - (map (fun p : nat * Bvector lambda => [i, a0]<-2 p; (i, a0)) - (zip (allNatsLt (length l0)) l0)) - = - (map (fun p => p) - (zip (allNatsLt (length l0)) l0)) - ). - eapply map_ext. - intuition. - rewrite H3. - clear H3. - rewrite map_id. - rewrite zip_combine_eq. - rewrite combine_split. - simpl. - rewrite combine_length. - rewrite allNatsLt_length. - rewrite Min.min_idempotent. - trivial. - eapply allNatsLt_length. - - intuition. - - eapply arrayLookupList_pred. - intuition. - destruct x0. - eapply in_combine_r in H2. - repeat intac. - subst. - simpl. - destruct x0. - unfold numberedMap. - rewrite map_length. - eapply In_zip in H3. - intuition. - repeat intac. - subst. - rewrite zip_length. - rewrite allNatsLt_length. - rewrite plus_0_r. - unfold maxMatch in *. - eapply arrayLookupList_pred. - eauto. - simpl. - omega. - eapply allNatsLt_length. - - simpl. - omega. - - Qed. - - Theorem ESPADA_TSetCorr_G13_G14_equiv : - Pr[ESPADA_TSetCorr_G13 maxMatches F A1] == Pr[ESPADA_TSetCorr_G14 maxMatches F A1]. - - rewrite ESPADA_TSetCorr_G13_1_equiv. - rewrite ESPADA_TSetCorr_G13_1_2_equiv. - rewrite ESPADA_TSetCorr_G13_2_3_equiv. - apply ESPADA_TSetCorr_G13_3_G14_equiv. - - Qed. - - Theorem ESPADA_TSetCorr_G14_G15_equiv : - Pr[ESPADA_TSetCorr_G14 maxMatches F A1] == Pr[ESPADA_TSetCorr_G15 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G14, ESPADA_TSetCorr_G15. - - comp_skip. - comp_simp. - comp_skip. - - case_eq (labelCollision_9 _ maxMatches F x); intuition. - simpl. - intuition. - - rewrite evalDist_ret_0. - rewrite evalDist_ret_0; - intuition. - - simpl. - intuition. - eapply negb_true_iff in H2. - eapply eqb_false_iff in H2. - - eapply map_ne_same_ex in H2. - destruct H2. - intuition. - eapply H4. - clear H4. - unfold ESPADA_TSetRetrieve_Corr_14. - eapply arrayLookupList_pred'. - intuition. - repeat rewrite <- zip_combine_eq in H2. - eapply list_pred_zip_in in H2. - Focus 2. - eapply list_pred_map_r. - eapply list_pred_eq. - simpl in *. - exdest. - intuition. - subst. - trivial. - - intuition. - eapply arrayLookup_None_not_In in H2. - intuition. - rewrite combine_split. - simpl. - trivial. - - rewrite map_length. - trivial. - - unfold eq_dec. - intuition. - eapply (EqDec_dec _). - - Qed. - - - Require Import PRF. - - - Definition ESPADA_TSetCorr_G15_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allLabels <-$ - (allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ret (map (fun x => map (F x) (allNatsLt (S maxMatches))) allTags)); - ret (hasDups _ (fst (split (flatten allLabels)))). - - Theorem ESPADA_TSetCorr_G15_1_equiv : - Pr[ESPADA_TSetCorr_G15 maxMatches F A1] == Pr[ESPADA_TSetCorr_G15_1]. - - unfold ESPADA_TSetCorr_G15, ESPADA_TSetCorr_G15_1. - - comp_skip. - comp_simp. - inline_first. - comp_skip. - comp_simp. - unfold labelCollision_9. - rewrite getAllLabels_8_1_equiv. - unfold getAllLabels_8_1. - - eapply evalDist_ret_eq. - f_equal. - rewrite fst_split_flatten_eq. - f_equal. - rewrite map_map. - eapply map_ext. - intuition. - rewrite fst_split_map_eq. - reflexivity. - - Qed. - - Theorem ESPADA_TSetCorr_G15_1_G16_equiv : - Pr[ESPADA_TSetCorr_G15_1] == Pr[ESPADA_TSetCorr_G16 maxMatches F A1]. - - unfold ESPADA_TSetCorr_G15_1, ESPADA_TSetCorr_G16. - unfold PRFI_A1, PRFI_A2. - inline_first. - comp_skip. - comp_simp. - comp_skip. - - eapply comp_spec_eq_impl_eq. - eapply comp_spec_eq_trans. - Focus 2. - eapply comp_spec_eq_symm. - eapply compMap_seq_map . - cbv beta. - simpl. - - eapply compMap_map_fission_eq. - intuition. - eapply comp_spec_consequence. - eapply comp_spec_eq_refl. - intuition; subst. - trivial. - Qed. - - Theorem ESPADA_TSetCorr_G15_G16_equiv : - Pr[ESPADA_TSetCorr_G15 maxMatches F A1] == - Pr[ESPADA_TSetCorr_G16 maxMatches F A1]. - - rewrite ESPADA_TSetCorr_G15_1_equiv. - eapply ESPADA_TSetCorr_G15_1_G16_equiv. - - Qed. - - Theorem ESPADA_TSetCorr_G16_G17_equiv : - | Pr[ESPADA_TSetCorr_G16 maxMatches F A1] - Pr[(ESPADA_TSetCorr_G17 bigB maxMatches A1)] | - == - PRF_NAI_Advantage (Rnd lambda) (RndF_R lambda bigB) F _ _ (PRFI_A1 maxMatches A1) (@PRFI_A2 lambda). - - unfold ESPADA_TSetCorr_G16, ESPADA_TSetCorr_G17, PRF_NAI_Advantage. - reflexivity. - - Qed. - - Theorem ESPADA_TSetCorr_G17_G18_equiv : - Pr[ESPADA_TSetCorr_G17 bigB maxMatches A1] == Pr[ESPADA_TSetCorr_G18 bigB maxMatches A1]. - - unfold ESPADA_TSetCorr_G17, ESPADA_TSetCorr_G18 . - unfold PRFI_A1, PRFI_A2. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply compMap_eq. - eapply list_pred_eq_in. - intuition. - intuition. - subst. - eapply in_map_iff in H3. - destruct H3. - intuition. - subst. - clear H0. - symmetry. - rewrite <- evalDist_right_ident. - eapply comp_spec_eq_impl_eq. - comp_skip. - unfold oracleMap. - eapply comp_spec_eq_trans_l. - eapply compMap_fold_equiv. - unfold compMap_fold. - specialize (compFold_spec); intuition. - eapply (@compFold_spec _ _ _ (fun ls a b => a = fst b /\ NoDup ls /\ forall x, In x ls -> arrayLookup _ (snd b) x = None)). - intuition. - eapply allNatsLt_NoDup. - intuition. - subst. - inversion H1; clear H1; subst. - comp_simp. - unfold RndR_func, randomFunc. - simpl in *. - rewrite H5. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - simpl. - case_eq (eqb x2 a); intuition. - rewrite eqb_leibniz in H9. - subst. - intuition. - intuition. - - comp_simp. - intuition; simpl in *; subst. - eapply comp_spec_eq_refl. - Qed. - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - forall t l, - In (t, l) (getSupport A1) -> - (length (combineKeywords (removeDups (list_EqDec bool_EqDec) l) (toW t)) <= maxKeywords)%nat. - - - Definition ESPADA_TSetCorr_G18_1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) lsD; - allQs' <- allNatsLt (maxKeywords - (length allQs)); - lsD' <- map (fun _ => allNatsLt (S maxMatches)) allQs'; - allLabels' <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) lsD'; - ret (hasDups _ (fst (split (flatten allLabels))) || hasDups _ (fst (split (flatten allLabels')))). - - Theorem ESPADA_TSetCorr_G18_1_equiv : - Pr[ ESPADA_TSetCorr_G18 bigB maxMatches A1] <= Pr[ESPADA_TSetCorr_G18_1 ]. - - unfold ESPADA_TSetCorr_G18, ESPADA_TSetCorr_G18_1 . - comp_skip. - comp_simp. - comp_skip. - comp_irr_r. - eapply compMap_wf. - intuition. - eapply compMap_wf. - intuition. - unfold RndF_R. - wftac. - eapply comp_spec_impl_le. - eapply comp_spec_ret; intuition. - - Qed. - - Definition ESPADA_TSetCorr_G18_2 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allQs' <- allNatsLt (maxKeywords - (length allQs)); - lsD' <- map (fun _ => allNatsLt (S maxMatches)) allQs'; - allLabels <-$ - (allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) lsD; - allLabels' <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) lsD'; - ret (allLabels ++ allLabels')); - ret (hasDups _ (fst (split (flatten allLabels)))). - - Theorem ESPADA_TSetCorr_G18_1_2_equiv : - Pr[ ESPADA_TSetCorr_G18_1] <= Pr[ESPADA_TSetCorr_G18_2 ]. - - unfold ESPADA_TSetCorr_G18_1, ESPADA_TSetCorr_G18_2. - comp_skip. - comp_simp. - inline_first. - comp_skip. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_impl_le. - eapply comp_spec_ret. - intuition. - - eapply hasDups_true_not_NoDup. - intuition. - eapply orb_true_iff in H2. - intuition. - eapply hasDups_true_not_NoDup in H4. - intuition. - eapply H4. - clear H4. - - rewrite fst_split_flatten_eq. - rewrite fst_split_flatten_eq in H3. - rewrite map_app in H3. - rewrite flatten_app in H3. - eapply NoDup_app_l. - eauto. - - eapply hasDups_true_not_NoDup in H4. - eapply H4. - clear H4. - rewrite fst_split_flatten_eq. - rewrite fst_split_flatten_eq in H3. - rewrite map_app in H3. - rewrite flatten_app in H3. - eapply NoDup_app in H3. - intuition. - Qed. - - Definition ESPADA_TSetCorr_G18_3 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allQs' <- allNatsLt (maxKeywords - (length allQs)); - lsD' <- map (fun _ => allNatsLt (S maxMatches)) allQs'; - allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) (lsD ++ lsD'); - ret (hasDups _ (fst (split (flatten allLabels)))). - - Theorem ESPADA_TSetCorr_G18_2_3_equiv : - Pr[ESPADA_TSetCorr_G18_2] == Pr[ESPADA_TSetCorr_G18_3]. - - unfold ESPADA_TSetCorr_G18_2, ESPADA_TSetCorr_G18_3. - comp_skip. - comp_simp. - comp_skip. - symmetry. - apply compMap_app. - - Qed. - - Definition ESPADA_TSetCorr_G18_4 := - allQs <- allNatsLt maxKeywords; - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => (RndF_R lambda bigB)) lsD) (lsD); - ret (hasDups _ (fst (split (flatten allLabels)))). - - Theorem ESPADA_TSetCorr_G18_3_4_equiv : - Pr[ESPADA_TSetCorr_G18_3] == Pr[ESPADA_TSetCorr_G18_4]. - - unfold ESPADA_TSetCorr_G18_3, ESPADA_TSetCorr_G18_4. - comp_irr_l. - comp_simp. - comp_skip. - - eapply compMap_eq. - - eapply list_pred_I_in. - rewrite app_length. - repeat rewrite map_length. - rewrite allNatsLt_length. - - specialize (maxKeywords_correct _ _ H). - unfold Blist in *. - - remember ((@length (list bool) - (combineKeywords - (@removeDups (list bool) (@list_EqDec bool bool_EqDec) l) - (@toW lambda t)))). - rewrite le_plus_minus_r. - symmetry. - eapply allNatsLt_length. - trivial. - - intuition. - intuition. - eapply in_app_or in H1. - intuition. - repeat intac. - subst. - reflexivity. - repeat intac. - subst. - reflexivity. - Qed. - - Definition ESPADA_TSetCorr_G18_5 := - allQs <- allNatsLt maxKeywords; - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => b <-$ RndNat bigB; l <-$ {0, 1}^lambda; ret (b, l)) lsD) (lsD); - ret (hasDups _ (flatten allLabels)). - - Theorem ESPADA_TSetCorr_G18_4_5_equiv : - Pr[ESPADA_TSetCorr_G18_4] == Pr[ESPADA_TSetCorr_G18_5]. - - unfold ESPADA_TSetCorr_G18_4, ESPADA_TSetCorr_G18_5. - comp_simp. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - - assert ( - comp_spec (fun a b => fst a = b) (RndF_R lambda bigB) - (b1 <-$ [ 0 .. bigB); l <-$ { 0 , 1 }^lambda; ret (b1, l)) - ). - unfold RndF_R. - - assert (Bvector lambda). - eapply (oneVector lambda). - assert (Bvector (S lambda)). - eapply (oneVector (S lambda)). - - comp_skip. - comp_skip. - comp_irr_l. - eapply comp_spec_ret; intuition. - eauto. - - eapply comp_spec_ret. - - symmetry. - erewrite list_pred_fst_split_flatten_eq_l. - reflexivity. - eauto. - Qed. - - Definition ESPADA_TSetCorr_G18_6 := - allQs <- allNatsLt maxKeywords; - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ - (x <-$ compMap _ (fun lsD => compMap _ (fun _ => b <-$ RndNat bigB; l <-$ {0, 1}^lambda; ret (b, l)) lsD) lsD; - ret (flatten x)); - ret (hasDups _ allLabels). - - Theorem ESPADA_TSetCorr_G18_5_6_equiv : - Pr[ESPADA_TSetCorr_G18_5] == Pr[ESPADA_TSetCorr_G18_6]. - - unfold ESPADA_TSetCorr_G18_5. - unfold ESPADA_TSetCorr_G18_6. - - comp_simp. - inline_first. - comp_skip. - comp_simp. - reflexivity. - - Qed. - - Definition ESPADA_TSetCorr_G18_7 := - allQs <- allNatsLt maxKeywords; - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ compMap _ (fun _ => b <-$ RndNat bigB; l <-$ {0, 1}^lambda; ret (b, l)) (flatten lsD); - ret (hasDups _ allLabels). - - Theorem ESPADA_TSetCorr_G18_6_7_equiv : - Pr[ESPADA_TSetCorr_G18_6] == Pr[ESPADA_TSetCorr_G18_7]. - - unfold ESPADA_TSetCorr_G18_6. - unfold ESPADA_TSetCorr_G18_7. - - comp_simp. - comp_skip. - - symmetry. - rewrite <- evalDist_right_ident. - symmetry. - eapply comp_spec_eq_impl_eq. - comp_skip. - eapply compMap_flatten. - simpl in *. - subst. - eapply comp_spec_eq_refl. - Qed. - - Theorem ESPADA_TSetCorr_G18_7_G19_equiv : - Pr[ESPADA_TSetCorr_G18_7] == Pr[ESPADA_TSetCorr_G19 lambda bigB maxMatches maxKeywords]. - - unfold ESPADA_TSetCorr_G18_7, ESPADA_TSetCorr_G19. - - comp_simp. - comp_skip. - eapply compMap_eq. - eapply list_pred_I. - - rewrite length_flatten. - rewrite fold_left_map_eq. - rewrite allNatsLt_length. - - rewrite fold_add_const_mult. - rewrite plus_0_r. - repeat rewrite allNatsLt_length. - trivial. - - intuition. - - Qed. - - Theorem ESPADA_TSetCorr_G18_G19_equiv : - Pr[ESPADA_TSetCorr_G18 bigB maxMatches A1] <= Pr[ESPADA_TSetCorr_G19 lambda bigB maxMatches maxKeywords]. - - rewrite ESPADA_TSetCorr_G18_1_equiv. - rewrite ESPADA_TSetCorr_G18_1_2_equiv. - rewrite ESPADA_TSetCorr_G18_2_3_equiv. - rewrite ESPADA_TSetCorr_G18_3_4_equiv. - rewrite ESPADA_TSetCorr_G18_4_5_equiv. - rewrite ESPADA_TSetCorr_G18_5_6_equiv. - rewrite ESPADA_TSetCorr_G18_6_7_equiv. - rewrite ESPADA_TSetCorr_G18_7_G19_equiv. - reflexivity. - Qed. - - Theorem ESPADA_TSetCorr_G19_G20_equiv : - Pr[ESPADA_TSetCorr_G19 lambda bigB maxMatches maxKeywords] <= - Pr[ESPADA_TSetCorr_G20 lambda maxMatches maxKeywords]. - - unfold ESPADA_TSetCorr_G19, ESPADA_TSetCorr_G20. - - eapply comp_spec_impl_le. - comp_skip. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - comp_irr_l. - eapply comp_spec_eq_trans_r. - Focus 2. - eapply comp_spec_right_ident. - comp_skip. - apply (oneVector lambda). - apply (oneVector lambda). - assert ( - comp_spec (fun x y => snd x = y) (ret (a, b0)) (ret b0) - ). - eapply comp_spec_ret. - trivial. - eauto. - - eapply comp_spec_ret. - intuition. - - eapply list_pred_snd_split_eq_l in H1. - subst. - eapply hasDups_true_not_NoDup. - intuition. - eapply hasDups_true_not_NoDup in H2. - eapply H2. - - eapply NoDup_snd_split_if. - trivial. - - Qed. - - Theorem ESPADA_TSetCorr_G20_small : - Pr[ESPADA_TSetCorr_G20 lambda maxMatches maxKeywords] <= - (maxKeywords * S maxMatches) ^ 2 / 2 ^ lambda. - - unfold ESPADA_TSetCorr_G20. - rewrite dupProb. - rewrite allNatsLt_length. - reflexivity. - Qed. - - Theorem ESPADA_TSetCorr_once_IPRF : - Pr[gInst ESPADA_TSetCorr_G1 ] <= (maxKeywords * (S maxMatches))^2 / 2 ^ lambda - + - PRF_NAI_Advantage ({ 0 , 1 }^lambda) (RndF_R lambda bigB) F nat_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) (PRFI_A1 maxMatches A1) (@PRFI_A2 lambda) - + - - PRF.PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar - (list_EqDec bool_EqDec) (Bvector_EqDec lambda) - (ESPADA_TSet_PRF_A1 - (pair_EqDec - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (z <-$ A1; [t, l]<-2 z; ret (t, l, (t, l)))) - (ESPADA_TSet_PRF_A2 bigB bigS F - (fun (s : (T lambda * list Blist)) - (p : option (TSet lambda) * list (Bvector lambda)) => - [t, q0] <-2 s; - [tSet_opt, tags]<-2 p; - tSet <- match tSet_opt with - | Some x => x - | None => nil - end; - ret ((if tSet_opt then true else false) && negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q0) - arrayLookupList (list_EqDec bool_EqDec) t x))))) - - . - - eapply leRat_trans. - eapply ratDistance_le_sum. - eapply ESPADA_TSetCorr_G1_G3_close. - eapply ratAdd_leRat_compat; [ idtac | reflexivity]. - - rewrite ESPADA_TSetCorr_G3_G4_equiv. - rewrite ESPADA_TSetCorr_G4_G5_equiv. - rewrite ESPADA_TSetCorr_G5_le_G6. - rewrite ESPADA_TSetCorr_G6_G7_equiv. - rewrite ESPADA_TSetCorr_G7_G8_equiv. - - rewrite ESPADA_TSetCorr_G8_le_G9. - rewrite ESPADA_TSetCorr_G9_G10_equiv. - rewrite ESPADA_TSetCorr_G10_G11_equiv. - rewrite ESPADA_TSetCorr_G11_G12_equiv. - rewrite ESPADA_TSetCorr_G12_G13_equiv. - rewrite ESPADA_TSetCorr_G13_G14_equiv. - rewrite ESPADA_TSetCorr_G14_G15_equiv. - rewrite ESPADA_TSetCorr_G15_G16_equiv. - - eapply leRat_trans. - eapply ratDistance_le_sum. - eapply eqRat_impl_leRat. - eapply ESPADA_TSetCorr_G16_G17_equiv. - eapply ratAdd_leRat_compat; [idtac | reflexivity]. - rewrite ESPADA_TSetCorr_G17_G18_equiv. - rewrite ESPADA_TSetCorr_G18_G19_equiv. - rewrite ESPADA_TSetCorr_G19_G20_equiv. - eapply ESPADA_TSetCorr_G20_small. - - Qed. - - Variable PRF_NA_Adv : Rat. - - Hypothesis PRF_NA_Adv_correct_1 : - PRF.PRF_NA_Advantage ({ 0 , 1 }^lambda) ({ 0 , 1 }^lambda) F_bar - (list_EqDec bool_EqDec) (Bvector_EqDec lambda) - (ESPADA_TSet_PRF_A1 - (pair_EqDec - (list_EqDec - (pair_EqDec (list_EqDec bool_EqDec) - (list_EqDec (Bvector_EqDec lambda)))) - (list_EqDec (list_EqDec bool_EqDec))) - (z <-$ A1; [t, l]<-2 z; ret (t, l, (t, l)))) - (ESPADA_TSet_PRF_A2 bigB bigS F - (fun (s : (T lambda * list Blist)) - (p : option (TSet lambda) * list (Bvector lambda)) => - [t, q0] <-2 s; - [tSet_opt, tags]<-2 p; - tSet <- match tSet_opt with - | Some x => x - | None => nil - end; - ret ((if tSet_opt then true else false) && negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q0) - arrayLookupList (list_EqDec bool_EqDec) t x))))) - <= PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_2: - forall i, - PRF_NA_Advantage ({ 0 , 1 }^lambda) (RndF_R lambda bigB) F _ _ - (Hybrid.B1 nil _ _ (PRFI_A1 maxMatches A1) i) - (Hybrid.B2 - (fun lsD => k <-$ {0, 1}^lambda; ret (map (F k) lsD)) - (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func (@RndF_range lambda bigB) _) nil lsD; ret lsR) _ - (PRFI_A2 lambda)) - <= PRF_NA_Adv. - - Theorem ESPADA_TSetCorr_once : - Pr[gInst ESPADA_TSetCorr_G1 ] <= - (maxKeywords * (S maxMatches))^2 / 2 ^ lambda - + - (maxKeywords / 1) * PRF_NA_Adv - + - - PRF_NA_Adv. - - rewrite ESPADA_TSetCorr_once_IPRF. - eapply ratAdd_leRat_compat; intuition. - eapply ratAdd_leRat_compat; intuition. - eapply PRF_NA_impl_NAI. - - intuition. - unfold PRFI_A1 in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - rewrite map_length. - eauto. - unfold RndF_R. - wftac. - wftac. - - eauto. - - Qed. - - -End ESPADA_TSet_Once_correct. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Correct_Once_Games.v b/src/ESPADA/ESPADA_TSet_Correct_Once_Games.v deleted file mode 100644 index 553ef6c..0000000 --- a/src/ESPADA/ESPADA_TSet_Correct_Once_Games.v +++ /dev/null @@ -1,615 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import TSet. -Require Import ESPADA_TSet. -Require Import CompFold. -Require Import Array. -Require Import ESPADA_TSet_Once_Games. -Require Import HasDups. -Require Import PRF. - -Local Open Scope list_scope. - - -Section ESPADA_TSet_Once_correct. - - Variable lambda : posnat. - Variable tSize bigB bigS maxMatches : posnat. - - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Variable A1 : Comp ((T lambda * list Blist)). - Hypothesis A1_wf : well_formed_comp A1. - - Variable maxKeywords : nat. - - Definition initFree := initFree bigB bigS. - Definition ESPADA_TSetRetrieve := ESPADA_TSetRetrieve lambda F maxMatches. - - (* Inline and simplify *) - Definition ESPADA_TSetCorr_G1 := - [t, l] <-$2 A1; - [tSet_opt, k_T] <-$2 ESPADA_TSetSetup_trial bigB bigS F F_bar t; - tSet <- getTSet tSet_opt; - tags <- map (F_bar k_T) l; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in l)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((if (tSet_opt) then true else false) && bad). - - Definition ESPADA_TSetSetup_tLoop_6 := - @ESPADA_TSetSetup_tLoop_6 lambda. - - (* Use part of the correctness proof to replace tags with random values. *) - Definition ESPADA_TSetCorr_G3 := - [t, q] <-$2 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((if loopRes then true else false) && bad). - - (* Simplify. We don't need to rebuild the list of queries and tags. *) - Definition ESPADA_TSetCorr_G4 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet loopRes; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((if loopRes then true else false) && bad). - - (* Put everything into the T-Set (and the list used to create it) *) - - Definition TSetCorr_5_Record := (Bvector lambda * ((Blist * Bvector lambda * nat * nat * Bvector lambda * nat * Bvector (S lambda)) * (Bvector (S lambda))))%type. - - Definition TSetCorr_5 := list (list (option TSetCorr_5_Record)). - - Definition tSetUpdate_Corr_5(tSet : TSetCorr_5) (b j : nat) - (r : TSetCorr_5_Record) := - let bucket := nth b tSet nil in - let bucket' := listReplace bucket j (Some r) None in - listReplace tSet b bucket' nil. - - - Definition ESPADA_TSetSetup_tLoop_Corr_5 (acc : TSetCorr_5 * FreeList) - (e : Blist * Bvector lambda * nat * nat * Vector.t bool lambda * nat * Bvector lambda * - Vector.t bool (S lambda)) := - [tSet, free]<-2 acc; - (let '(q, tag, i, len, s_i, b, bigL, bigK) := e in - free_b <- nth b free nil; - opt_j <-$ RndListElem.rndListElem nat_EqDec free_b; - match opt_j with - | Some j => - free0 <- - listReplace free b (removeFirst (EqDec_dec nat_EqDec) free_b j) nil; - bet <- (if eq_nat_dec (S i) len then false else true); - newRecord <- (bigL, ((q, tag, i, len, s_i, b, bigK), (Vector.cons bool bet lambda s_i xor bigK))); - tSet0 <- tSetUpdate_Corr_5 tSet b j newRecord; ret Some (tSet0, free0) - | None => ret None - end). - - Definition getTSet_Corr_5 (o : option (TSetCorr_5 * FreeList)) := - match o with - | None => nil - | Some p => fst p - end. - - Definition ESPADA_TSetRetrieve_tLoop_Corr_5 (tSet : TSetCorr_5) stag i := - [b, L, K] <-3 F stag i; - B <- nth b tSet nil; - t <- arrayLookupOpt _ B L; - match t with - | None => None - | Some p => - [_, u] <-2 p; - v <- u xor K; - bet <- Vector.hd v; - s <- Vector.tl v; - Some (s, bet) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_5 (tSet : TSetCorr_5) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_5 tSet stag i) with - | Some (v, bet) => - v :: (if (bet) then (ESPADA_TSetRetrieve_h_Corr_5 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_5 tSet stag := - ESPADA_TSetRetrieve_h_Corr_5 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G5 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_Corr_5) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_Corr_5 loopRes; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_5 tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((if loopRes then true else false) && bad). - - (* Assume no collisions in labels within a bucket *) - Fixpoint labelIn_6(ls : list (option TSetCorr_5_Record)) l := - match ls with - | nil => false - | o :: ls' => - match o with - | None => labelIn_6 ls' l - | Some (l', _) => - if (eqb l l') then true else (labelIn_6 ls' l) - end - end. - - - Fixpoint bucketLabelCollision_6 (ls : list (option TSetCorr_5_Record)) := - match ls with - | nil => false - | o :: ls' => - match o with - | None => bucketLabelCollision_6 ls' - | Some (l, _) => - labelIn_6 ls' l || bucketLabelCollision_6 ls' - end - end. - - - Fixpoint labelCollision_6 (ls : list (list (option TSetCorr_5_Record))) := - match ls with - | nil => false - | b :: ls' => - bucketLabelCollision_6 b || labelCollision_6 ls' - end. - - - Definition ESPADA_TSetCorr_G6 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_Corr_5) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_Corr_5 loopRes; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_5 tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((if loopRes then true else false) && (labelCollision_6 (tSet) || bad)). - - (* Don't randomize order during TSet setup. Always succeed (possibly by overfilling a bucket)*) - Fixpoint labelIn_7(ls : list (TSetCorr_5_Record)) l := - match ls with - | nil => false - | (l', _) :: ls' => - if (eqb l l') then true else (labelIn_7 ls' l) - end. - - Fixpoint bucketLabelCollision_7 (ls : list ( TSetCorr_5_Record)) := - match ls with - | nil => false - | (l, _) :: ls' => - labelIn_7 ls' l || bucketLabelCollision_7 ls' - end. - - - Fixpoint labelCollision_7 (ls : list (list ( TSetCorr_5_Record))) := - match ls with - | nil => false - | b :: ls' => - bucketLabelCollision_7 b || labelCollision_7 ls' - end. - - Definition TSetCorr_7 := list (list (TSetCorr_5_Record)). - - Definition ESPADA_TSetSetup_tLoop_Corr_7 (tSet : TSetCorr_7) - (e : Blist * Bvector lambda * nat * nat * Vector.t bool lambda * nat * Bvector lambda * - Vector.t bool (S lambda)) := - (let '(q, tag, i, len, s_i, b, bigL, bigK) := e in - bet <- (if eq_nat_dec (S i) len then false else true); - newRecord <- (bigL, ((q, tag, i, len, s_i, b, bigK), (Vector.cons bool bet lambda s_i xor bigK))); - newBucket <- (nth b tSet nil) ++ (newRecord :: nil); - listReplace tSet b newBucket nil - ). - - Definition ESPADA_TSetRetrieve_tLoop_Corr_7 (tSet : TSetCorr_7) stag i := - [b, L, K] <-3 F stag i; - B <- nth b tSet nil; - t <- arrayLookup _ B L; - match t with - | None => None - | Some p => - [_, u] <-2 p; - v <- u xor K; - bet <- Vector.hd v; - s <- Vector.tl v; - Some (s, bet) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_7 (tSet : TSetCorr_7) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_7 tSet stag i) with - | Some (v, bet) => - v :: (if (bet) then (ESPADA_TSetRetrieve_h_Corr_7 tSet stag (S i) fuel') else nil) - | None => nil - end - - end. - - Definition ESPADA_TSetRetrieve_Corr_7 tSet stag := - ESPADA_TSetRetrieve_h_Corr_7 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G7 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- fold_left ( ESPADA_TSetSetup_tLoop_Corr_7) (flatten ts_recsList) nil; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_7 tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret (labelCollision_7 ( tSet) || bad). - - (* flatten the list *) - Definition TSetCorr_8_Record := ((nat * Bvector lambda) * ((Blist * Bvector lambda * nat * nat * Bvector lambda * nat * Bvector (S lambda)) * (Bvector (S lambda))))%type. - - Fixpoint labelIn_8 (Z : Set)(ls : list ((nat * Bvector lambda) * Z)) l := - match ls with - | nil => false - | (l', _) :: ls' => - if (eqb l l') then true else (labelIn_8 ls' l) - end. - - Fixpoint labelCollision_8 (Z : Set)(ls : list ((nat * Bvector lambda) * Z)) := - match ls with - | nil => false - | (l, _) :: ls' => - labelIn_8 ls' l || labelCollision_8 ls' - end. - - Definition TSetCorr_8 := list (TSetCorr_8_Record). - - Definition ESPADA_TSetSetup_tLoop_Corr_8 - (e : Blist * Bvector lambda * nat * nat * Vector.t bool lambda * nat * Bvector lambda * - Vector.t bool (S lambda)) := - (let '(q, tag, i, len, s_i, b, bigL, bigK) := e in - ((b, bigL), ((q, tag, i, len, s_i, b, bigK), (Vector.cons bool (if eq_nat_dec (S i) len then false else true) lambda s_i xor bigK))) - ). - - Definition ESPADA_TSetRetrieve_tLoop_Corr_8 (tSet : TSetCorr_8) stag i := - [b, L, K] <-3 F stag i; - t <- arrayLookup _ tSet (b, L); - match t with - | None => None - | Some p => - [_, u] <-2 p; - v <- u xor K; - bet <- Vector.hd v; - s <- Vector.tl v; - Some (s, bet) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_8 (tSet : TSetCorr_8) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_8 tSet stag i) with - | Some (v, bet) => - v :: (if (bet) then (ESPADA_TSetRetrieve_h_Corr_8 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_8 tSet stag := - ESPADA_TSetRetrieve_h_Corr_8 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G8 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- map ( ESPADA_TSetSetup_tLoop_Corr_8) (flatten ts_recsList) ; - tags <- firstn (length q) allTags; - bad <- negb - (eqb (foreach (x in tags)ESPADA_TSetRetrieve_Corr_8 tSet x) - (foreach (x in q)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret (labelCollision_8 ( tSet) || bad). - - (* assume no collisions in tags or F output *) - Definition labelCollision_9_tag t := - hasDups _ (map (F t) (allNatsLt maxMatches)). - - Fixpoint getAllLabels tags := - match tags with - | nil => nil - | t :: tags' => - (map (fun i => fst (F t i)) (allNatsLt (S maxMatches))) ++ (getAllLabels tags') - end. - - Definition labelCollision_9 tags := - hasDups _ (getAllLabels tags). - - Definition ESPADA_TSetCorr_G9 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- map ( ESPADA_TSetSetup_tLoop_Corr_8) (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allTags)ESPADA_TSetRetrieve_Corr_8 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - (* use the stored values and counts rather than the encryptions *) - Definition ESPADA_TSetRetrieve_tLoop_Corr_10 (tSet : TSetCorr_8) stag i := - [b, L, K] <-3 F stag i; - t <- arrayLookup _ tSet (b, L); - match t with - | None => None - | Some (q, tag, i, len, s_i, b, bigK, u) => - Some (s_i, if (eq_nat_dec (S i) len) then false else true) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_10 (tSet : TSetCorr_8) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_10 tSet stag i) with - | Some (v, b) => - v :: (if b then (ESPADA_TSetRetrieve_h_Corr_10 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_10 tSet stag := - ESPADA_TSetRetrieve_h_Corr_10 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G10 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (q, tag, i, len, s_i, b, bigL, bigK)) ls) (zip (zip allQs allTags) ts); - tSet <- map ( ESPADA_TSetSetup_tLoop_Corr_8) (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allTags)ESPADA_TSetRetrieve_Corr_10 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - - (* Simplify *) - Definition TSetCorr_11_Record := ((nat * Bvector lambda) * ((Blist * Bvector lambda * nat * nat * Bvector lambda)))%type. - - Definition TSetCorr_11 := list TSetCorr_11_Record. - - Definition ESPADA_TSetRetrieve_tLoop_Corr_11 (tSet : TSetCorr_11) stag i := - [b, L, K] <-3 F stag i; - t <- arrayLookup _ tSet (b, L); - match t with - | None => None - | Some (q, tag, i, len, s_i) => - Some (s_i, if (eq_nat_dec (S i) len) then false else true) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_11 (tSet : TSetCorr_11) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_11 tSet stag i) with - | Some (v, b) => - v :: (if b then (ESPADA_TSetRetrieve_h_Corr_11 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_11 tSet stag := - ESPADA_TSetRetrieve_h_Corr_11 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G11 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; ((b, bigL), (q, tag, i, len, s_i))) ls) (zip (zip allQs allTags) ts); - tSet <- (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allTags)ESPADA_TSetRetrieve_Corr_11 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - (* lookup by tag and index---this is the same as long as there are no collisions. *) - Definition TSetCorr_12_Record := ((Bvector lambda * nat) * ((Blist * nat * Bvector lambda)))%type. - - Definition TSetCorr_12 := list TSetCorr_12_Record. - - Definition ESPADA_TSetRetrieve_tLoop_Corr_12 (tSet : TSetCorr_12) stag i := - t <- arrayLookup _ tSet (stag, i); - match t with - | None => None - | Some (q, len, s_i) => - Some (s_i, if (eq_nat_dec (S i) len) then false else true) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_12 (tSet : TSetCorr_12) stag i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_12 tSet stag i) with - | Some (v, b) => - v :: (if b then (ESPADA_TSetRetrieve_h_Corr_12 tSet stag (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_12 tSet stag := - ESPADA_TSetRetrieve_h_Corr_12 tSet stag O maxMatches. - - Definition ESPADA_TSetCorr_G12 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; ((tag, i), (q, len, s_i))) ls) (zip (zip allQs allTags) ts); - tSet <- (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allTags)ESPADA_TSetRetrieve_Corr_12 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - (* Look up using the keyword and index *) - Definition TSetCorr_13_Record := ((Blist * nat) * (nat * (Bvector lambda)))%type. - - Definition TSetCorr_13 := list TSetCorr_13_Record. - - Definition ESPADA_TSetRetrieve_tLoop_Corr_13 (tSet : TSetCorr_13) q i := - t <- arrayLookup _ tSet (q, i); - match t with - | None => None - | Some (len, s_i) => - Some (s_i, if (eq_nat_dec (S i) len) then false else true) - end. - - Fixpoint ESPADA_TSetRetrieve_h_Corr_13 (tSet : TSetCorr_13) q i (fuel : nat) := - match fuel with - | O => nil - | S fuel' => - match (ESPADA_TSetRetrieve_tLoop_Corr_13 tSet q i) with - | Some (v, b) => - v :: (if b then (ESPADA_TSetRetrieve_h_Corr_13 tSet q (S i) fuel') else nil) - | None => nil - end - end. - - Definition ESPADA_TSetRetrieve_Corr_13 tSet q := - ESPADA_TSetRetrieve_h_Corr_13 tSet q O maxMatches. - - Definition ESPADA_TSetCorr_G13 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [q, tag, ls] <-3 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; ((q, i), (len, s_i))) ls) (zip (zip allQs allTags) ts); - tSet <- (flatten ts_recsList) ; - bad <- negb - (eqb (foreach (x in allQs)ESPADA_TSetRetrieve_Corr_13 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - (* Now we can replace the retrieve function with a simple array lookup. *) - Definition TSetCorr_14_Record := (Bvector lambda)%type. - - Definition TSetCorr_14 := list (Blist * list TSetCorr_14_Record). - - Definition ESPADA_TSetRetrieve_Corr_14 (tSet : TSetCorr_14) (q : Blist) := - arrayLookupList _ tSet q. - - Definition ESPADA_TSetCorr_G14 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - tSet <- combine allQs ts; - bad <- negb - (eqb (foreach (x in allQs)ESPADA_TSetRetrieve_Corr_14 tSet x) - (foreach (x in allQs)arrayLookupList (list_EqDec bool_EqDec) t x)); - ret ((labelCollision_9 allTags) || bad). - - (* "bad" is always false. Simplify *) - Definition ESPADA_TSetCorr_G15 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ret ((labelCollision_9 allTags)). - - - (* factor out the calls to the PRF and prepare to apply the PRF definition. *) - Definition PRFI_A1 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - ret (map (fun _ => allNatsLt (S maxMatches)) allQs, tt). - - Definition PRFI_A2 (s : unit) (l : list (list (nat * Bvector lambda * Bvector (S lambda)))) := - ret (hasDups _ (fst (split (flatten l)))). - - Definition ESPADA_TSetCorr_G16 := - [lsD, s_A] <-$2 PRFI_A1; - allLabels <-$ compMap _ (fun ls => k <-$ {0, 1}^lambda; ret (map (F k) ls)) lsD; - PRFI_A2 s_A allLabels. - - (* Replace PRF with random function *) - Definition RndF_R := - n <-$ RndNat bigB; - l <-$ {0, 1}^lambda; - k <-$ {0, 1}^(S lambda); - ret (n, l, k). - - Definition ESPADA_TSetCorr_G17 := - [lsD, s_A] <-$2 PRFI_A1; - allLabels <-$ compMap _ (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func RndF_R _) nil lsD; ret lsR) lsD; - PRFI_A2 s_A allLabels. - - (* No duplicates in input---replace random function output with random values *) - Definition ESPADA_TSetCorr_G18 := - [t, q] <-$2 A1; - q <- removeDups _ q; - allQs <- combineKeywords q (toW t); - lsD <- map (fun _ => allNatsLt (S maxMatches)) allQs; - allLabels <-$ compMap _ (fun lsD => compMap _ (fun _ => RndF_R) lsD) lsD; - ret (hasDups _ (fst (split (flatten allLabels)))). - - (* get labels for the maximum number of keywords *) - Definition ESPADA_TSetCorr_G19 := - allLabels <-$ compMap _ (fun _ => b <-$ RndNat bigB; l <-$ {0, 1}^lambda; ret (b, l)) (allNatsLt (maxKeywords * (S maxMatches))); - ret (hasDups _ allLabels). - - (* The chance of duplicates in labels is sufficiently small---ignore buckets. *) - Definition ESPADA_TSetCorr_G20 := - allLabels <-$ compMap _ (fun _ => {0, 1}^lambda) (allNatsLt (maxKeywords * (S maxMatches))); - ret (hasDups _ allLabels). - - -End ESPADA_TSet_Once_correct. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Once.v b/src/ESPADA/ESPADA_TSet_Once.v deleted file mode 100644 index 273b093..0000000 --- a/src/ESPADA/ESPADA_TSet_Once.v +++ /dev/null @@ -1,3519 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import FCF. -Require Import Array. -Require Import CompFold. -Require Import RndNat. -Require Import RndListElem. -Require Import PRF. - -Require Import TSet. -Require Import ESPADA_TSet_Once_Games. - -Local Open Scope array_scope. -Local Open Scope list_scope. - -Section ESPADA_TSet_once. - - Variable lambda : posnat. - Variable tSize bigB bigS : posnat. - Variable bucketSize : nat -> (posnat * posnat). - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Definition TSet := (@TSet lambda). - Definition initFree := (@initFree bigB bigS). - - Hypothesis F_b_correct : - forall k b i v1 v2, - (b, v1, v2) = F k i -> - b < bigB. - - - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (T lambda * list Blist * A_State). - Variable A2 : A_State -> (option TSet) * list (Bvector lambda) -> Comp bool. - - - Notation "'ESPADA_TSetSetup_Sim_wLoop'" := (ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop F) (at level 90). - - - Notation "'gInst' g" := (@g lambda bigB bigS F F_bar A_State A1 A2) (at level 90). - - - (* Step 1 : Inline and simplify. *) - Theorem ESPADA_TSet_1_equiv: - Pr[TSetReal _ (@ESPADA_TSetSetup_once lambda bigB bigS F F_bar) (@ESPADA_TSetGetTag lambda F_bar) A1 A2] == Pr[ gInst ESPADA_TSet_1 ]. - - unfold ESPADA_TSet_1, TSetReal, ESPADA_TSetSetup, ESPADA_TSetSetup_trial. - comp_skip. - comp_simp. - inline_first. - comp_skip. - inline_first. - comp_skip. - reflexivity. - eapply trans. - eapply evalDist_seq; - intuition. - eapply comp_spec_eq_impl_eq. - eapply compMap_map. - eapply eqRat_refl. - comp_simp. - eapply eqRat_refl. - Qed. - - (* Step 2: pull out the stag computation and answer lookup. We switch to the simulator setup loop because it happens to work here. *) - Theorem ESPADA_TSet_2_equiv : - Pr[gInst ESPADA_TSet_1] == Pr[gInst ESPADA_TSet_2]. - - unfold ESPADA_TSet_1, ESPADA_TSet_2. - - do 3 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply (@compFold_eq). - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_map_l_eq. - eapply list_pred_map_l_eq. - eapply list_pred_map_l_eq. - - intuition. - intuition. - destruct H4. - intuition. - - destruct a2; simpl in *; subst. - unfold ESPADA_TSetSetup_wLoop, ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop. - comp_simp. - unfold foldBody_option. - eapply comp_spec_eq_refl. - - Qed. - - - (* Step 3: initialize the TSet using all keywords -- including those that are in the the queries but not in the structure. *) - - Fixpoint lookupIndex (A : Set)(eqd : eq_dec A)(ls : list A)(a : A) def := - match ls with - | nil => def - | a' :: ls' => - if (eqd a a') then O else S (lookupIndex eqd ls' a def) - end. - - (* Step 3_1 : Start by initializing the absent keywords at then end -- this will be a no-op *) - Definition ESPADA_TSet_3_1 := - [t, q, s_A] <-$3 A1; - k_T <-$ {0, 1} ^ lambda; - stags <- map (F_bar k_T) (toW t); - ts <- map (arrayLookupList _ t) (toW t); - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip stags ts); - absent_qs <- (removePresent (EqDec_dec _) (toW t) (removeDups _ q)); - q_tags <- map (F_bar k_T) absent_qs ; - q_ts <- map (arrayLookupList _ t) absent_qs; - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) loopRes (zip q_tags q_ts); - tSet <- getTSet_opt loopRes; - tags <- map (F_bar k_T) q; - A2 s_A (tSet, tags). - - - (* Step 3_2 : Combine the two loops into one *) - Definition ESPADA_TSet_3_2 := - [t, q, s_A] <-$3 A1; - absent_qs <- (removePresent (EqDec_dec _) (toW t) (removeDups _ q)); - k_T <-$ {0, 1} ^ lambda; - stags <- map (F_bar k_T) (toW t); - q_tags <- map (F_bar k_T) absent_qs ; - ts <- map (arrayLookupList _ t) (toW t); - q_ts <- map (arrayLookupList _ t) absent_qs; - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) ((zip stags ts) ++ (zip q_tags q_ts)); - tSet <- getTSet_opt loopRes; - tags <- map (F_bar k_T) q; - A2 s_A (tSet, tags). - - Ltac hypInv := - try (match goal with - | [H: Some _ = Some _ |-_ ] => inversion H; clear H; subst - end); try pairInv. - - Theorem ESPADA_TSet_3_1_equiv : - Pr[gInst ESPADA_TSet_2] == Pr[ESPADA_TSet_3_1]. - - unfold ESPADA_TSet_2, ESPADA_TSet_3_1. - - comp_skip; comp_simp. - comp_skip. - comp_skip. - reflexivity. - - comp_irr_r. - - eapply compFold_wf; intuition. - unfold foldBody_option. - destruct a0; wftac. - unfold ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop;wftac. - destruct p. - unfold setLet. - eapply compFold_wf; intuition. - unfold foldBody_option. - destruct a0; wftac. - unfold ESPADA_TSetSetup_tLoop; wftac. - - destruct p. - unfold ESPADA_TSet.ESPADA_TSetSetup_tLoop . - unfold maybeBindComp. - wftac. - eapply rndListElem_wf. - - intuition. - destruct b4; wftac. - - apply compFold_nop in H2; intuition. - - subst. - intuition. - - apply In_zip in H3. - intuition. - apply in_map_iff in H5. - apply in_map_iff in H6. - destruct H5. - destruct H6. - intuition. - subst. - - unfold foldBody_option in *. - destruct x0. - - unfold ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop in *. - destruct p. - unfold setLet in *. - - assert ((@arrayLookupList Blist (Bvector (posnatToNat lambda)) - (@list_EqDec bool bool_EqDec) t x3) = nil). - unfold arrayLookupList. - - case_eq (t # x3); intuition. - exfalso. - eapply removePresent_not_in. - eauto. - unfold toW. - eapply in_removeDups. - - eapply arrayLookup_Some_In_unzip; eauto. - - rewrite H3 in H4. - simpl in H4. - intuition. - simp_in_support. - trivial. - Qed. - - Theorem ESPADA_TSet_3_2_equiv : - Pr[ESPADA_TSet_3_1] == Pr[ESPADA_TSet_3_2]. - - unfold ESPADA_TSet_3_1, ESPADA_TSet_3_2. - - do 2 (comp_skip; comp_simp); - rewrite <- evalDist_assoc. - - symmetry. - comp_skip. - eapply compFold_app; intuition. - - intuition. - Qed. - - Theorem ESPADA_TSet_3_equiv : - Pr[ESPADA_TSet_3_2] == Pr[gInst ESPADA_TSet_3]. - - unfold ESPADA_TSet_3_2, ESPADA_TSet_3. - - comp_skip; comp_simp. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply comp_spec_eq_impl_eq. - eapply (compFold_eq). - - rewrite zip_app. - repeat rewrite <- map_app. - unfold combineKeywords. - eapply list_pred_eq. - - repeat rewrite map_length. - trivial. - repeat rewrite map_length. - trivial. - - intuition. - subst. - eapply comp_spec_eq_refl. - - rewrite map_map. - symmetry. - - erewrite (map_ext_in _ _ (F_bar x)); - intuition. - - erewrite nth_map_In. - rewrite nth_lookupIndex . - trivial. - - Lemma combineKeywords_in : - forall ls1 ls2 a, - (In a ls1 \/ In a ls2) -> - In a (combineKeywords ls1 ls2). - - intros. - destruct (in_dec (EqDec_dec _) a ls1). - clear H. - unfold combineKeywords. - eapply in_or_app. - intuition. - - intuition. - unfold combineKeywords. - eapply in_or_app. - right. - eapply removePresent_In; intuition. - - Qed. - - eapply combineKeywords_in; intuition. - right. - eapply in_removeDups. - trivial. - - eapply lookupIndex_lt_length. - eapply combineKeywords_in ; intuition. - right. - eapply in_removeDups. - trivial. - - Grab Existential Variables. - apply nil. - - Qed. - - - Require Import OTP. - Notation "'gInst1' g" := (g lambda bigB bigS F A_State A1 A2) (at level 90). - - Theorem ESPADA_TSet_15_equiv : - Pr[gInst1 ESPADA_TSet_14] == Pr[gInst1 ESPADA_TSet_15]. - - unfold ESPADA_TSet_14, ESPADA_TSet_15. - - do 4 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply comp_fold_ext; intuition. - - unfold foldBody_option. - destruct a0. - - unfold ESPADA_TSetSetup_tLoop_14, ESPADA_TSetSetup_tLoop_15. - comp_simp. - comp_skip. - comp_skip. - eapply comp_spec_eq_symm. - eapply comp_spec_seq. - apply None. - eapply None. - eapply xor_OTP; intuition. - intuition. - subst. - comp_skip. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_refl. - Qed. - - - Theorem ESPADA_TSet_16_equiv : - Pr[gInst1 ESPADA_TSet_15] == Pr[gInst1 ESPADA_TSet_16]. - - unfold ESPADA_TSet_15, ESPADA_TSet_16. - - do 4 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_I. - - eapply length_flatten_eq. - - eapply list_pred_impl. - eapply list_pred_map_r. - - eapply list_pred_map_l. - eapply list_pred_map_l_eq. - - intuition. - destruct H2. - intuition. - destruct H3. - intuition. - subst. - - rewrite numberedMap_length. - trivial. - - intuition. - - unfold foldBody_option. - destruct acc. - unfold ESPADA_TSetSetup_tLoop_15, randomTSetEntry. - comp_simp. - comp_swap_l. - comp_skip. - comp_swap_l. - comp_skip. - comp_skip. - comp_skip. - destruct b5. - comp_simp. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_refl. - Qed. - - Theorem toW_NoDup : - forall (t : T lambda), - NoDup (toW t). - - intuition. - apply removeDups_NoDup. - Qed. - - Theorem ESPADA_TSet_17_equiv : - Pr[gInst1 ESPADA_TSet_16] == Pr[gInst1 ESPADA_TSet_17]. - - unfold ESPADA_TSet_16, ESPADA_TSet_17. - do 4 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_I. - rewrite allNatsLt_length. - - unfold combineKeywords. - - rewrite skipn_app. - - repeat rewrite length_flatten. - repeat rewrite fold_left_map_eq. - Require Import Permutation. - - rewrite fold_left_add_removePresent. - f_equal. - - symmetry. - - rewrite (fold_add_nat_filter_partition (fun a => if (in_dec (EqDec_dec _) a (toW t)) then true else false)). - - rewrite plus_comm. - rewrite fold_add_nat_0. - rewrite plus_0_l. - intuition. - - intuition. - apply filter_In in H2. - intuition. - destruct (@in_dec Blist (@EqDec_dec Blist (@list_EqDec bool bool_EqDec)) - a0 (toW t)). - simpl in *. - discriminate. - - unfold toW in *. - unfold arrayLookupList. - - rewrite notInArrayLookupNone. - simpl. - trivial. - intuition. - eapply n. - - eapply in_removeDups. - trivial. - - eapply toW_NoDup. - eapply removeDups_NoDup. - - intuition. - eapply comp_spec_eq_refl. - - Theorem eqRat_eq : - forall x y, - x = y -> - x == y. - - intuition; subst; intuition. - - Qed. - - eapply eqRat_eq. - f_equal. - f_equal. - f_equal. - f_equal. - eapply map_ext_in. - intuition. - unfold combineKeywords. - - Theorem lookupIndex_app : - forall (A : Set)(eqd : eq_dec A)(l1 l2 : list A)(a : A) n, - In a l1 -> - CompFold.lookupIndex eqd (l1 ++ l2) a n = - CompFold.lookupIndex eqd l1 a n. - - induction l1; intuition; simpl in *. - intuition. - - intuition; subst. - destruct (eqd a0 a0); intuition. - - - destruct (eqd a0 a); intuition. - - - Qed. - - eapply lookupIndex_app . - eapply in_removeDups. - trivial. - - Qed. - - Theorem ESPADA_TSet_18_equiv : - Pr[gInst1 ESPADA_TSet_17] == Pr[gInst1 ESPADA_TSet_18]. - - unfold ESPADA_TSet_17, ESPADA_TSet_18. - do 3 (comp_skip; comp_simp). - - symmetry. - - eapply (@compFold_flatten _ _ _ _ eq). - - unfold flatten_prep. - unfold numberedMap. - eapply list_pred_eq. - - intuition. - unfold foldBody_option. - destruct init. - unfold ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop. - comp_simp. - unfold foldBody_option. - eapply comp_spec_impl_eq. - eapply comp_spec_consequence. - eapply (@compFold_eq _ _ (fun (p1 : nat * Bvector lambda) p2 => let (i1, s_i1) := p1 in let '(i2, len, tag, s_i2) := p2 in - i1 = i2 /\ s_i1 = s_i2 /\ tag = d /\ len = (length lsa))). - - eapply list_pred_map_l_inv in H2. - - eapply list_pred_impl. - eapply H2. - - intuition. - subst. - intuition. - intuition. - - destruct acc. - - Theorem ESPADA_TSetSetup_tLoop_4_equiv : - forall tag len acc i s_i x, - evalDist (ESPADA_TSetSetup_tLoop F tag len acc (i, s_i)) x == - evalDist (ESPADA_TSetSetup_tLoop_4 F acc (i, len, tag, s_i)) x. - - intuition. - unfold ESPADA_TSetSetup_tLoop, ESPADA_TSetSetup_tLoop_4. - - - unfold ESPADA_TSet.ESPADA_TSetSetup_tLoop. - remember (F tag i) as v. - destruct v. - comp_simp. - comp_skip. - destruct x0. - comp_simp. - destruct (eq_nat_dec (S i) len). - rewrite e. - rewrite eqb_refl. - unfold negb. - eapply eqRat_refl. - case_eq (eqb (S i) len); intuition. - rewrite eqb_leibniz in H1. - intuition. - unfold negb. - eapply eqRat_refl. - - intuition. - Qed. - - destruct a2. - destruct p0. - destruct p0. - intuition; subst. - eapply eq_impl_comp_spec_eq. - eapply (ESPADA_TSetSetup_tLoop_4_equiv). - eapply comp_spec_eq_refl. - - intuition; subst; intuition. - - specialize (compFold_foldBodyOption_None ); intuition. - unfold foldBody_option in *. - rewrite H3. - intuition. - Qed. - - - Theorem removeDups_pred : - forall (A B : Set) (f : A -> B)(f' : B -> A) (eqda : EqDec A)(eqdb : EqDec B)(lsa : list A)(lsb : list B), - list_pred (fun a b => b = f a) lsa lsb -> - (forall a, In a lsa -> f' (f a) = a) -> - list_pred (fun a b => b = f a) (removeDups _ lsa) (removeDups _ lsb). - - induction 1; intuition; simpl in *. - econstructor. - - destruct (in_dec (EqDec_dec eqda) a1 ls1). - destruct (in_dec (EqDec_dec eqdb) a2 ls2). - eauto. - - subst. - exfalso. - eapply n. - - Theorem list_pred_In_exists : - forall (A B : Set) P (lsa : list A)(lsb : list B), - list_pred P lsa lsb -> - forall a, In a lsa -> - exists b, In b lsb /\ P a b. - - induction 1; intuition; simpl in *. - intuition. - intuition; subst. - econstructor. - intuition. - - edestruct IHlist_pred. - eauto. - intuition. - exists x. - intuition. - - Qed. - - edestruct list_pred_In_exists. - eapply IHlist_pred. - intuition. - eapply in_removeDups. - eauto. - intuition. - subst. - eapply removeDups_in. - intuition. - - subst. - destruct (in_dec (EqDec_dec eqdb) (f a1) ls2). - exfalso. - eapply n. - rewrite <- (@H1 a1). - eapply list_pred_symm in IHlist_pred. - edestruct list_pred_In_exists. - eauto. - eapply in_removeDups. - eauto. - intuition. - rewrite H3. - rewrite H1. - eapply removeDups_in. - trivial. - - right. - eapply removeDups_in. - trivial. - - intuition. - intuition. - - econstructor; eauto. - - Qed. - - - Theorem removeDups_NoDup_eq : - forall (A : Set)(eqd : EqDec A)(ls : list A), - NoDup ls -> - removeDups eqd ls = ls. - - induction 1; intuition; simpl in *. - destruct (in_dec (EqDec_dec eqd) x l); intuition. - f_equal. - intuition. - Qed. - - Definition ESPADA_TSet_19_1 := - [t, q, s_A]<-$3 A1; - eqPattern <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ eqPattern); - q_ts <-$ ret map (arrayLookupList _ t) (removeDups _ q); - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry bigB acc)) loopRes (allNatsLt (minus (bigN t) (length (flatten q_ts)))); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_19_1_equiv : - Pr[gInst1 ESPADA_TSet_18] == Pr[ESPADA_TSet_19_1]. - - unfold ESPADA_TSet_18, ESPADA_TSet_19_1. - - comp_skip. - comp_simp. - - comp_skip. - eapply compMap_eq. - - eapply (@removeDups_pred _ _ (fun a => CompFold.lookupIndex _ (removeDups _ l) a O) (fun b => nth b (removeDups _ l) nil)). - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - reflexivity. - intuition. - eapply nth_lookupIndex. - eapply in_removeDups. - trivial. - - intuition. - comp_simp. - reflexivity. - Qed. - - Definition ESPADA_TSet_19_2 := - [t, q, s_A]<-$3 A1; - eqPattern <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ eqPattern); - - uniqueTs <-$ ret map (fun n => nth (firstIndexOf (EqDec_dec _ ) eqPattern n O) (map (arrayLookupList _ t) q) nil) (removeDups _ eqPattern); - - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip qTags uniqueTs); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry bigB acc)) loopRes (allNatsLt (minus (bigN t) (length (flatten uniqueTs)))); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_19_2_equiv : - Pr[ESPADA_TSet_19_1] == Pr[ESPADA_TSet_19_2]. - - unfold ESPADA_TSet_19_1, ESPADA_TSet_19_2. - comp_skip. - comp_simp. - eapply evalDist_seq_eq. - intuition. - intros. - comp_skip. - eapply evalDist_ret_eq. - - eapply map_eq. - - Theorem list_pred_In : - forall (A B : Set) (P : A -> B -> Prop) (lsa : list A)(lsb : list B), - list_pred (fun a b => In a lsa -> In b lsb -> P a b) lsa lsb -> - list_pred P lsa lsb. - - induction lsa; inversion 1; intuition; simpl in *. - econstructor. - subst. - econstructor. - eapply H2; intuition. - eapply IHlsa. - eapply list_pred_impl. - eapply H4. - intuition. - - Qed. - - eapply list_pred_In. - eapply list_pred_impl. - - - eapply (@removeDups_pred _ _ (fun a => CompFold.lookupIndex _ (removeDups _ l) a O) (fun b => nth b (removeDups _ l) nil)). - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - intuition; subst. - reflexivity. - intuition. - eapply nth_lookupIndex. - eapply in_removeDups. - trivial. - - intuition. - subst. - - Theorem firstIndexOf_map_1_1 : - forall (B A : Set) (f : B -> A) (eqda : eq_dec A)(eqdb : eq_dec B)(ls : list B) (a : B), - (forall a2, In a2 ls -> f a = f a2 -> a = a2) -> - firstIndexOf eqda (map f ls) (f a) O = - firstIndexOf eqdb ls a O. - - induction ls; intuition; simpl in *. - destruct (eqdb a0 a); subst. - - destruct (eqda (f a) (f a)); intuition. - - destruct ( eqda (f a0) (f a)). - exfalso. - eapply n. - eapply H; intuition. - - f_equal. - eauto. - Qed. - - rewrite (@firstIndexOf_map_1_1 _ _ (fun a0 => CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) a0 0) (EqDec_dec _) (EqDec_dec _)). - - erewrite map_nth_in. - rewrite nth_lookupIndex. - trivial. - - eapply removeDups_in. - trivial. - - eapply firstIndexOf_in_lt. - eapply removeDups_in. - trivial. - - intuition. - - Theorem lookupIndex_1_1 : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(a1 a2 : A) n, - In a1 ls -> - In a2 ls -> - lookupIndex eqd ls a1 n = lookupIndex eqd ls a2 n -> - a1 = a2. - - induction ls; intuition; simpl in *. - intuition. - - intuition; subst; intuition. - - destruct (eqd a2 a1 ); intuition. - destruct (eqd a1 a1); intuition. - discriminate. - - destruct ( eqd a1 a2); intuition. - destruct (eqd a2 a2); intuition. - discriminate. - - destruct (eqd a1 a); - destruct ( eqd a2 a); subst; intuition; try discriminate. - - eapply IHls; intuition. - - inversion H1; clear H1; subst. - eauto. - - Qed. - - eapply lookupIndex_1_1; eauto. - eapply in_removeDups. - trivial. - - Grab Existential Variables. - eapply nil. - Qed. - - Theorem ESPADA_TSet_19_equiv' : - Pr[ESPADA_TSet_19_2] == Pr[gInst1 ESPADA_TSet_19]. - - unfold ESPADA_TSet_19_2, ESPADA_TSet_19. - comp_skip; comp_simp. - comp_skip; comp_simp. - reflexivity. - reflexivity. - - Qed. - - Theorem ESPADA_TSet_19_equiv : - Pr[gInst1 ESPADA_TSet_18] == Pr[gInst1 ESPADA_TSet_19]. - - rewrite ESPADA_TSet_19_1_equiv. - rewrite ESPADA_TSet_19_2_equiv. - eapply ESPADA_TSet_19_equiv'. - - Qed. - - Theorem Ideal_equiv : - Pr[gInst1 ESPADA_TSet_19] == Pr[TSetIdeal (@L_T lambda) A1 A2 (ESPADA_TSet_Sim_once lambda bigB bigS F)]. - - unfold ESPADA_TSet_19, ESPADA_TSet_Sim, ESPADA_TSet_Sim_trial, TSetIdeal, L_T. - - comp_skip; comp_simp. - unfold ESPADA_TSet_Sim_once, ESPADA_TSet_Sim_trial. - inline_first. comp_skip. - inline_first. comp_skip. - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_eq. - intuition; subst. - eapply comp_spec_eq_refl. - - inline_first. comp_skip. - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_eq. - intuition; subst. - eapply comp_spec_eq_refl. - - comp_simp. - reflexivity. - Qed. - - - - Theorem ESPADA_TSet_4_equiv : - Pr[gInst ESPADA_TSet_3] == Pr[gInst ESPADA_TSet_4]. - - unfold ESPADA_TSet_3, ESPADA_TSet_4. - - do 2 (comp_skip; comp_simp). - comp_skip. - eapply compFold_flatten. - unfold flatten_prep. - unfold numberedMap. - eapply list_pred_eq. - - intuition. - unfold ESPADA_TSet_Once_Games.ESPADA_TSetSetup_Sim_wLoop. - - unfold foldBody_option. - destruct init. - comp_simp. - eapply comp_spec_eq_impl_eq. - eapply (@compFold_eq _ _ (fun (p1 : nat * Bvector lambda) p2 => let (i1, s_i1) := p1 in let '(i2, len, tag, s_i2) := p2 in - i1 = i2 /\ s_i1 = s_i2 /\ tag = d /\ len = (length lsa))). - - eapply list_pred_map_l_inv in H1. - eapply list_pred_impl. - eapply H1. - - intuition. - subst. - intuition. - intuition. - - destruct a2. - destruct p. - destruct p. - intuition; subst. - - destruct acc. - eapply eq_impl_comp_spec_eq. - eapply ESPADA_TSetSetup_tLoop_4_equiv. - eapply comp_spec_eq_refl. - - specialize (compFold_foldBodyOption_None ); intuition. - unfold foldBody_option in *. - rewrite H2. - intuition. - - Qed. - - Theorem ESPADA_TSet_6_equiv : - Pr[gInst ESPADA_TSet_4] == Pr[gInst ESPADA_TSet_6]. - - unfold ESPADA_TSet_4, ESPADA_TSet_6. - - do 3 (comp_skip; comp_simp). - eapply comp_spec_eq_impl_eq. - eapply (@compFold_eq _ _ (fun (p1 : nat * nat * Bvector lambda * Bvector lambda) (p2 : nat * nat * Bvector lambda * nat * Bvector lambda * Bvector (S lambda)) => - let '(i1, len1, tag1, s_i1) := p1 in - let '(i2, len2, s_i2, b2, bigL2, bigK2) := p2 in - i1 = i2 /\ len1 = len2 /\ s_i1 = s_i2 /\ (b2, bigL2, bigK2) = F tag1 i1)). - - eapply list_pred_flatten_both. - - eapply list_pred_impl. - eapply list_pred_map_both. - eapply list_pred_eq. - intuition. - destruct H1. - destruct H1. - intuition. - subst. - comp_simp. - unfold numberedMap. - eapply list_pred_impl. - eapply list_pred_map_both. - eapply list_pred_eq. - intuition. - destruct H1. - destruct H1. - intuition. - subst. - destruct x2. - pairInv. - - remember (F b n) as v1. - destruct v1. - destruct p. - intuition. - - intuition. - destruct a2. - destruct p. - destruct p. - destruct p. - destruct p. - intuition; subst. - - unfold ESPADA_TSetSetup_tLoop_4, ESPADA_TSetSetup_tLoop_6. - unfold foldBody_option. - destruct acc. - comp_simp. - pairInv. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_refl. - Qed. - - Require Import Permutation. - - - - Definition correctFreeList (free : FreeList) := - (length free = bigB) /\ (forall ls, In ls free -> NoDup ls /\ (forall j, In j ls -> j < bigS)). - - Theorem ESPADA_TSet_7_equiv : - Pr[gInst ESPADA_TSet_6] == Pr[gInst ESPADA_TSet_7]. - - unfold ESPADA_TSet_6, ESPADA_TSet_7. - - comp_skip; comp_simp. - inline_first. - comp_skip. - comp_simp. - - comp_skip. - eapply (compFold_perm - (fun acc => - match acc with - | None => True - | Some (tSet, free) => - correctFreeList free - end)); intuition. - - eapply Permutation_flatten. - eapply Permutation_map. - - repeat rewrite zip_map_eq. - eapply Permutation_map. - - Theorem combineKeywords_Permutation : - forall ls1 ls2, - NoDup ls1 -> - NoDup ls2 -> - Permutation (combineKeywords ls1 ls2) (combineKeywords ls2 ls1). - - intuition. - eapply NoDup_Permutation; intuition. - - - Theorem combineKeywords_NoDup : - forall (ls1 ls2 : list Blist), - NoDup ls1 -> - NoDup ls2 -> - NoDup (combineKeywords ls1 ls2). - - unfold combineKeywords; intuition. - eapply app_NoDup; intuition. - eapply removePresent_NoDup; intuition. - eapply removePresent_correct. - eapply H1. - eauto. - - eapply removePresent_correct. - eapply H2. - eauto. - Qed. - - - apply combineKeywords_NoDup; intuition. - apply combineKeywords_NoDup; intuition. - - - Theorem combineKeywords_in_iff : - forall a ls1 ls2, - In a (combineKeywords ls1 ls2) <-> (In a ls1 \/ In a ls2). - - unfold combineKeywords in *; intuition. - apply in_app_or in H; intuition. - right. - eapply removePresent_in_only_if; eauto. - - eapply in_or_app. - destruct (in_dec Blist_eq_dec a ls1); intuition. - right. - - eapply removePresent_correct2; intuition. - - Qed. - - eapply combineKeywords_in_iff in H1. - eapply combineKeywords_in_iff; intuition. - - eapply combineKeywords_in_iff in H1. - eapply combineKeywords_in_iff; intuition. - - Qed. - - eapply combineKeywords_Permutation. - - - eapply toW_NoDup. - eapply removeDups_NoDup. - - apply in_flatten in H1. - apply in_flatten in H2. - destruct H1. - destruct H2. - intuition. - apply in_map_iff in H4. - apply in_map_iff in H1. - destruct H4. - destruct H1. - intuition. - destruct x4. - destruct x5. - subst. - - unfold numberedMap in *. - apply in_map_iff in H5. - apply in_map_iff in H6. - destruct H5. - destruct H6. - intuition. - destruct x2. - destruct x3. - remember (F b12 n0) as v2. - remember (F b11 n) as v3. - destruct v2. - destruct p. - destruct v3. - destruct p. - repeat pairInv. - - unfold foldBody_option. - destruct b. - destruct p. - - Definition ESPADA_TSetSetup_tLoop_6' tSet (e : nat * nat * Bvector lambda * nat * Bvector lambda * Bvector (S lambda)) j := - let '(i, len, s_i, b, bigL, bigK) := e in - bet <- if (eq_nat_dec (S i) len) then false else true; - newRecord <- (bigL, BVxor _ (Vector.cons _ bet _ s_i) bigK); - tSetUpdate tSet b j newRecord. - - Lemma tLoop_6_pair_equiv : - forall tSet free len1 len2 i1 i2 s_i1 s_i2 b1 b2 bigL1 bigL2 bigK1 bigK2 x, - evalDist - (b' <-$ ESPADA_TSetSetup_tLoop_6 (tSet, free) (i1, len1, s_i1, b1, bigL1, bigK1); - match b' with - | Some b1 => ESPADA_TSetSetup_tLoop_6 b1 (i2, len2, s_i2, b2, bigL2, bigK2) - | None => ret None - end) x == - evalDist - (opt_j1_j2 <-$ - (free_b1 <- nth b1 free nil; - opt_j1 <-$ rndListElem _ (free_b1); - match opt_j1 with - | None => ret None - | Some j1 => - free <- listReplace free b1 (removeFirst (EqDec_dec _) free_b1 j1) nil; - free_b2 <- nth b2 free nil; - opt_j2 <-$ rndListElem _ (free_b2); - match opt_j2 with - | None => ret None - | Some j2 => - ret (Some (j1, j2)) - end - end); - match opt_j1_j2 with - | None => ret None - | Some (j1, j2) => - tSet <- ESPADA_TSetSetup_tLoop_6' tSet (i1, len1, s_i1, b1, bigL1, bigK1) j1; - tSet <- ESPADA_TSetSetup_tLoop_6' tSet (i2, len2, s_i2, b2, bigL2, bigK2) j2; - free <- listReplace free b1 (removeFirst (EqDec_dec _) (nth b1 free nil) j1) nil; - free <- listReplace free b2 (removeFirst (EqDec_dec _) (nth b2 free nil) j2) nil; - ret (Some (tSet, free)) - end) - x. - - intuition. - unfold ESPADA_TSetSetup_tLoop_6'. - unfold ESPADA_TSetSetup_tLoop_6. - - comp_simp. - inline_first. - comp_skip. - destruct x0. - inline_first. - comp_skip. - destruct x0. - comp_simp. - - eapply evalDist_ret_eq. - f_equal. - comp_simp. - intuition. - comp_simp. - intuition. - Qed. - - repeat rewrite tLoop_6_pair_equiv. - comp_simp. - (* case split on whether the buckets are the same *) - destruct (eq_nat_dec b2 b8). - subst. - - eapply (evalDist_iso (fun x => optSwap x) (fun x => optSwap x)); intuition. - apply optSwap_involutive. - apply optSwap_involutive. - - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - unfold optSwap. - rewrite <- rndListElem_support in *. - rewrite nth_listReplace_eq in *. - - destruct (eq_nat_dec n0 n); subst. - eapply getSupport_In_Seq. - rewrite <- rndListElem_support. - eauto. - - cbv beta iota. - eapply getSupport_In_Seq. - rewrite <- rndListElem_support. - rewrite nth_listReplace_eq in *. - eauto. - simpl. - intuition. - - eapply getSupport_In_Seq. - rewrite <- rndListElem_support. - eapply removeFirst_in_iff. - eauto. - cbv beta iota. - eapply getSupport_In_Seq. - rewrite <- rndListElem_support. - rewrite nth_listReplace_eq in *. - eapply removeFirst_in. - eauto. - intuition. - simpl. - intuition. - - repeat simp_in_support. - eapply getSupport_In_Seq. - eapply H2. - cbv beta iota. - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - repeat simp_in_support. - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - rewrite <- rndListElem_support in *. - rewrite nth_listReplace_eq in *. - - eapply comp_spec_impl_eq. - eapply comp_spec_seq; try eapply None. - eapply eq_impl_comp_spec; - try apply rndListElem_wf. - eapply rndListElem_uniform. - eapply H3. - eapply nth_In. - assert (length f = bigB). - eapply H3. - rewrite H1. - eapply F_b_correct; eauto. - - rewrite <- rndListElem_support. - eapply removeFirst_in_iff. - eauto. - rewrite <- rndListElem_support. - eauto. - - intuition. - - (* - - eapply H4. - - simpl in *. - intuition. - discriminate. - - simpl in *. - intuition. - discriminate. - - repeat simp_in_support. - destruct x1. - repeat simp_in_support. - destruct x1. - simp_in_support. - hypInv. - trivial. - simp_in_support. - discriminate. - simp_in_support. - discriminate. - - destruct r0. - destruct r1. - *) - - destruct a2. - destruct b. - repeat rewrite nth_listReplace_eq. - simpl. - destruct (eq_nat_dec n1 n0); subst. - assert (Some n2 = Some n). - intuition. - hypInv. - - eapply comp_spec_seq; try eapply None. - - eapply (comp_spec_iso - (fun x => - match x with - | None => None - | Some v => Some - (if (eq_nat_dec v n) then n0 else (if (eq_nat_dec v n0) then n else v)) - end) - (fun x => - match x with - | None => None - | Some v => Some - (if (eq_nat_dec v n) then n0 else (if (eq_nat_dec v n0) then n else v)) - end)); intuition. - - destruct x2. - destruct (eq_nat_dec n1 n); subst. - destruct (eq_nat_dec n0 n); subst; intuition. - destruct (eq_nat_dec n0 n0); subst; intuition. - destruct (eq_nat_dec n1 n0); subst. - destruct (eq_nat_dec n n); subst; intuition. - destruct (eq_nat_dec n1 n); subst. - intuition. - destruct (eq_nat_dec n1 n0); subst; intuition. - trivial. - - destruct x2. - destruct (eq_nat_dec n1 n); subst. - destruct (eq_nat_dec n0 n); subst; intuition. - destruct (eq_nat_dec n0 n0); subst; intuition. - destruct (eq_nat_dec n1 n0); subst. - destruct (eq_nat_dec n n); subst; intuition. - destruct (eq_nat_dec n1 n); subst. - intuition. - destruct (eq_nat_dec n1 n0); subst; intuition. - trivial. - - destruct x2. - destruct (eq_nat_dec n1 n); subst. - rewrite <- rndListElem_support in *. - exfalso. - eapply removeFirst_NoDup_not_in; eauto. - eapply nth_NoDup; intuition. - eapply H3; intuition. - - destruct (eq_nat_dec n1 n0); subst. - rewrite <- rndListElem_support in *. - eapply removeFirst_in. - trivial. - intuition. - - rewrite <- rndListElem_support in *. - eapply removeFirst_in. - eapply removeFirst_in_iff. - eauto. - intuition. - - rewrite <- rndListElem_support in *. - - rewrite rndListElem_support_None in *. - exfalso. - eapply (@in_nil nat). - rewrite <- H12. - eapply H4. - - destruct x2. - destruct (eq_nat_dec n1 n); subst. - - eapply rndListElem_uniform_gen. - eapply removeFirst_NoDup. - eapply H3. - eapply nth_In. - rewrite <- rndListElem_support in *. - eapply lt_le_trans. - eapply F_b_correct; eauto. - unfold correctFreeList in *. - intuition. - - eapply removeFirst_NoDup. - eapply H3. - eapply nth_In. - rewrite <- rndListElem_support in *. - eapply lt_le_trans. - eapply F_b_correct; eauto. - unfold correctFreeList in *. - intuition. - - Theorem removeFirst_length_In : - forall (A : Set)(eqd : eq_dec A)(ls : list A)(x : A), - In x ls -> - length (removeFirst eqd ls x) = pred (length ls). - - induction ls; intuition; simpl in *. - intuition; subst. - destruct (eqd x x); try discriminate; intuition. - - destruct (eqd x a); subst; intuition. - simpl. - erewrite IHls; intuition. - symmetry. - eapply S_pred. - destruct ls. - simpl in *. - intuition. - simpl in *. - assert (0 < S (length ls)). - omega. - eauto. - Qed. - - Show. - repeat erewrite removeFirst_length_In; eauto. - eapply removeFirst_in_iff. - eauto. - trivial. - - eapply removeFirst_in; intuition. - subst. - eapply removeFirst_NoDup_not_in; eauto. - eapply nth_NoDup; intuition. - eapply H3; intuition. - - - eapply rndListElem_uniform_gen. - unfold correctFreeList in *. - eapply removeFirst_NoDup. - eapply H3. - eapply nth_In. - intuition. - rewrite H13. - eapply F_b_correct; eauto. - - unfold correctFreeList in *. - eapply removeFirst_NoDup. - eapply H3. - eapply nth_In. - intuition. - rewrite H13. - eapply F_b_correct; eauto. - - repeat rewrite removeFirst_length_In; intuition. - eapply removeFirst_in_iff; eauto. - rewrite <- rndListElem_support in *. - destruct (eq_nat_dec n1 n0); subst. - exfalso. - eapply removeFirst_NoDup_not_in; eauto. - eapply nth_NoDup; intuition. - eapply H3; intuition. - eapply removeFirst_in. - eapply removeFirst_in_iff; eauto. - intuition. - - rewrite <- rndListElem_support in *. - eauto. - - (* contradiction. There are at least two elements in the list, so None is not in the support after we remove one. *) - rewrite <- rndListElem_support in *. - rewrite rndListElem_support_None in *. - exfalso. - eapply (@in_nil nat). - rewrite <- H12. - eapply removeFirst_in. - eapply H9. - intuition. - subst. - eapply removeFirst_NoDup_not_in; eauto. - eapply nth_NoDup; intuition. - eapply H3; intuition. - - intuition. - destruct a2. - destruct (eq_nat_dec n1 n); subst. - - eapply comp_spec_ret; intuition. - - destruct (eq_nat_dec n1 n0); subst. - eapply comp_spec_ret; intuition. - hypInv; intuition. - hypInv; intuition. - - eapply comp_spec_ret; intuition; - hypInv; intuition. - - subst. - eapply comp_spec_ret; intuition; - discriminate. - - destruct (eq_nat_dec n2 n); subst; intuition. - repeat hypInv. - intuition. - - comp_irr_l. - eapply rndListElem_wf. - comp_irr_r. - eapply rndListElem_wf. - destruct a2; - destruct b; - eapply comp_spec_ret; intuition; hypInv; intuition; try discriminate. - - rewrite <- rndListElem_support in *. - rewrite rndListElem_support_None in *. - rewrite H9 in H1. - simpl in *. - intuition. - - destruct b. - rewrite <- rndListElem_support in H9. - rewrite rndListElem_support_None in *. - rewrite H1 in H9. - simpl in *. - intuition. - - eapply comp_spec_ret; intuition. - simpl in *; discriminate. - discriminate. - - rewrite <- rndListElem_support in *. - simp_in_support. - comp_skip. - destruct x2. - comp_skip. - destruct x2. - repeat rewrite evalDist_ret_0; intuition. - discriminate. - simpl in *; discriminate. - simpl; intuition. - simpl; intuition. - - simp_in_support. - comp_skip. - destruct x2. - comp_skip. - destruct x2. - simpl; intuition. - simpl; intuition. - simpl; intuition. - - - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - unfold optSwap. - - destruct (eq_nat_dec n n0); subst. - (* impossible, means we got the same location twice when sampling (without replacement) from the free list for the bucket *) - rewrite <- rndListElem_support in *. - rewrite nth_listReplace_eq in H4. - exfalso. - eapply removeFirst_NoDup_not_in. - eapply H3. - eapply nth_In. - unfold correctFreeList in *. - intuition. - rewrite H1. - eapply F_b_correct. - eauto. - eauto. - - repeat rewrite nth_listReplace_eq. - eapply evalDist_ret_eq. - f_equal. - - Lemma ESPADA_TSetSetup_tLoop_6'_comm : - forall tSet j1 j2 - len1 i1 s_i1 b1 bigL1 bigK1 - len2 i2 s_i2 b2 bigL2 bigK2, - (b1 <> b2 \/ j1 <> j2) -> - ESPADA_TSetSetup_tLoop_6' (ESPADA_TSetSetup_tLoop_6' tSet (len1, i1, s_i1, b1, bigL1, bigK1) j1) (len2, i2, s_i2, b2, bigL2, bigK2) j2 = - ESPADA_TSetSetup_tLoop_6' (ESPADA_TSetSetup_tLoop_6' tSet (len2, i2, s_i2, b2, bigL2, bigK2) j2) (len1, i1, s_i1, b1, bigL1, bigK1) j1. - - - unfold ESPADA_TSetSetup_tLoop_6'; intros. - comp_simp. - - Lemma listReplace_comm : - forall (A : Set)(i1 i2 : nat)(ls : list A)(a1 a2 def : A), - i1 <> i2 -> - listReplace (listReplace ls i1 a1 def) i2 a2 def = - listReplace (listReplace ls i2 a2 def) i1 a1 def. - - induction i1; destruct i2; destruct ls; intuition; simpl in *. - f_equal. - eapply IHi1; omega. - - f_equal. - eapply IHi1; omega. - - Qed. - - Lemma tSetUpdate_comm : - forall tSet b1 b2 j1 j2 r1 r2, - (b1 <> b2 \/ j1 <> j2) -> - tSetUpdate (@tSetUpdate lambda tSet b2 j2 r2) b1 j1 r1 = - tSetUpdate (tSetUpdate tSet b1 j1 r1) b2 j2 r2. - - unfold tSetUpdate; intros. - destruct (eq_nat_dec b1 b2); subst. - intuition. - - repeat rewrite listReplace_twice. - - f_equal. - - repeat rewrite nth_listReplace_eq. - apply listReplace_comm ; intuition. - - clear H. - - repeat rewrite (@nth_listReplace_ne b1 b2); intuition. - repeat rewrite (@nth_listReplace_ne b2 b1); intuition. - - apply listReplace_comm; intuition. - - - Qed. - - eapply tSetUpdate_comm. - intuition. - Qed. - - f_equal. - - eapply ESPADA_TSetSetup_tLoop_6'_comm; intuition. - - repeat rewrite listReplace_twice. - f_equal. - - Lemma removeFirst_comm : - forall (A : Set)(eqd : eq_dec A)(ls : list A) (a1 a2 : A), - removeFirst eqd (removeFirst eqd ls a1) a2 = - removeFirst eqd (removeFirst eqd ls a2) a1. - - induction ls; intuition; simpl in *. - destruct (eqd a1 a); subst. - destruct (eqd a2 a); subst. - trivial. - simpl. - destruct (eqd a a); subst; intuition. - destruct (eqd a2 a); subst. - simpl. - destruct (eqd a a); intuition. - - simpl. - destruct (eqd a2 a); subst; intuition. - destruct (eqd a1 a); subst; intuition. - - f_equal. - eapply IHls. - - Qed. - - eapply removeFirst_comm. - - simp_in_support. - unfold optSwap. - intuition. - - simp_in_support. - simpl; intuition. - - (* Now we do the case in which the buckets are different. These computations are independent, so we can just swap things around. *) - eapply (evalDist_iso (fun x => optSwap x) (fun x => optSwap x)); intuition. - apply optSwap_involutive. - apply optSwap_involutive. - repeat simp_in_support. - destruct x3. - simp_in_support. - destruct x3. - repeat simp_in_support. - - unfold optSwap. - - eapply getSupport_In_Seq. - rewrite <- rndListElem_support in *. - - Theorem nth_listReplace_nil_ne : - forall (A : Set)(i2 i1 : nat) (a : A) def, - i1 <> i2 -> - nth i1 (listReplace nil i2 a def) def = - def. - - induction i2; intuition; simpl in *. - - destruct i1; intuition. - destruct i1; intuition. - - destruct i1; intuition. - Qed. - - rewrite nth_listReplace_ne in H4; intuition. - eauto. - comp_simp. - eapply getSupport_In_Seq. - rewrite nth_listReplace_ne; intuition. - eauto. - simpl. - intuition. - - simp_in_support. - rewrite nth_listReplace_ne in H4; intuition. - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - simp_in_support. - - edestruct (@rndListElem_support_exists nat). - eapply getSupport_In_Seq. - eauto. - destruct x2. - eapply getSupport_In_Seq. - rewrite nth_listReplace_ne; intuition. - eauto. - simpl. - intuition. - simpl. intuition. - - eapply (trans - _ (evalDist - (opt_j1 <-$ rndListElem nat_EqDec (nth b2 f nil); - opt_j2 <-$ rndListElem nat_EqDec (nth b8 f nil); - match opt_j1 with - | Some j1 => - match opt_j2 with - | Some j2 => ret (Some (j1, j2)) - | None => ret None - end - | None => ret None - end) (optSwap x2))). - - comp_skip. - destruct x3. - rewrite nth_listReplace_ne; intuition. - comp_irr_r. - eapply rndListElem_wf. - - eapply (trans - _ (evalDist - (opt_j2 <-$ rndListElem nat_EqDec (nth b8 f nil); - opt_j1 <-$ rndListElem nat_EqDec (nth b2 f nil); - match opt_j1 with - | Some j1 => - match opt_j2 with - | Some j2 => ret (Some (j1, j2)) - | None => ret None - end - | None => ret None - end) (optSwap x2))). - - comp_swap leftc. - comp_skip. - - comp_skip. - destruct x3. - rewrite nth_listReplace_ne; intuition. - comp_skip. - destruct x3; intuition. - - destruct x2. - destruct p. - unfold optSwap. - destruct (EqDec_dec (option_EqDec (pair_EqDec nat_EqDec nat_EqDec)) - (Some (n3, n2)) (Some (n1, n0))). - hypInv. - repeat rewrite evalDist_ret_1; intuition. - repeat rewrite evalDist_ret_0; intuition; hypInv; intuition. - - unfold optSwap. - repeat rewrite evalDist_ret_0; intuition; discriminate. - - destruct x2. - destruct p. - unfold optSwap. - - repeat rewrite evalDist_ret_0; intuition; discriminate. - - simpl. - intuition. - - comp_irr_l. - eapply rndListElem_wf. - destruct x3. - destruct x2. - destruct p. - unfold optSwap. - repeat rewrite evalDist_ret_0; intuition; discriminate. - simpl. - intuition. - - destruct x2. - destruct p. - unfold optSwap. - repeat rewrite evalDist_ret_0; intuition; discriminate. - - unfold optSwap. - intuition. - - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - destruct x3. - repeat simp_in_support. - unfold optSwap. - - repeat rewrite nth_listReplace_ne; intuition. - eapply evalDist_ret_eq. - f_equal. - f_equal. - - eapply ESPADA_TSetSetup_tLoop_6'_comm; intuition; eauto using F_b_correct. - rewrite listReplace_comm; intuition. - - simp_in_support. - unfold optSwap in *. - intuition. - simp_in_support. - unfold optSwap. - intuition. - - comp_simp. - intuition. - - unfold foldBody_option in H3. - destruct b5. - destruct p. - unfold ESPADA_TSetSetup_tLoop_6 in *. - simp_in_support. - destruct x1. - simp_in_support. - - unfold correctFreeList in *; intuition. - apply in_flatten in H1. - destruct H1. - intuition. - eapply in_map_iff in H2. - destruct H2. - intuition. - destruct x2. - unfold numberedMap in *. - subst. - eapply in_map_iff in H6. - destruct H6. - intuition. - destruct x1. - remember (F b5 n0) as v1. - destruct v1. - destruct p. - pairInv. - - rewrite listReplace_length; intuition. - rewrite H3. - eapply F_b_correct; eauto. - - apply in_flatten in H1. - destruct H1. - intuition. - apply in_map_iff in H6. - destruct H6. - intuition. - destruct x2. - unfold numberedMap in *. - subst. - apply in_map_iff in H7. - destruct H7. - intuition. - destruct x1. - remember (F b5 n0) as v1. - destruct v1. - destruct p. - pairInv. - - apply listReplace_in in H2; intuition. - eapply H5. - trivial. - subst. - - eapply removeFirst_NoDup. - eapply H5. - eapply nth_In. - rewrite H3. - - eapply F_b_correct. - eauto. - subst. - econstructor. - - - apply in_flatten in H1. - destruct H1. - intuition. - apply in_map_iff in H7. - destruct H7. - destruct x2. - intuition. - subst. - unfold numberedMap in *. - apply in_map_iff in H8. - destruct H8. - intuition. - destruct x1. - remember (F b5 n0) as v1. - destruct v1. - destruct p. - pairInv. - - apply listReplace_in in H2; intuition; subst. - eapply H5; eauto. - - eapply H5. - eapply nth_In. - rewrite H3. - eapply F_b_correct; eauto. - eapply removeFirst_in_iff. - eauto. - simpl in H6. - intuition. - - simp_in_support. - trivial. - - simp_in_support. - trivial. - - Theorem initFree_correct : - correctFreeList (initFree). - - unfold correctFreeList, initFree, ESPADA_TSet_Once_Games.initFree, ESPADA_TSet.initFree. - split. - rewrite map_length. - eapply allNatsLt_length. - - intros. - apply in_map_iff in H. - destruct H. - intuition; subst. - eapply map_NoDup. - eapply allNatsLt_NoDup. - intuition. - apply in_map_iff in H. - destruct H. - intuition. - subst. - apply allNatsLt_lt; intuition. - - Qed. - - apply initFree_correct. - - eapply eqRat_eq. - f_equal. - f_equal. - f_equal. - repeat rewrite map_map. - eapply map_ext_in. - intuition. - - repeat erewrite nth_map_In; - repeat erewrite nth_lookupIndex; intuition. - eapply combineKeywords_in. - left. - eapply in_removeDups. - trivial. - eapply combineKeywords_in. - right. - eapply in_removeDups. - trivial. - - eapply lookupIndex_lt_length. - eapply combineKeywords_in. - left. - eapply in_removeDups. - trivial. - - eapply lookupIndex_lt_length. - eapply combineKeywords_in. - right. - eapply in_removeDups. - trivial. - - - Grab Existential Variables. - apply nil. - apply nil. - - - Theorem firstn_combineKeywords : - forall ls1 ls2, - firstn (length ls1) (combineKeywords ls1 ls2) = ls1. - - intuition. - unfold combineKeywords. - eapply firstn_app_eq. - Qed. - - Qed. - - Definition ESPADA_TSet_PRF_A1 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - ts <- map (arrayLookupList _ t) allQs; - ret (allQs, (s_A, ts, allQs, q)). - - Definition ESPADA_TSet_PRF_A2 s_A1 allTags := - let '(s_A, ts, allQs, q) := s_A1 in - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun x => nth x allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Definition ESPADA_TSet_8_1 := - [allQs, s_A1] <-$2 ESPADA_TSet_PRF_A1; - allTags <-$ (k_T <-$ {0, 1} ^ lambda; ret (map (F_bar k_T) allQs)); - ESPADA_TSet_PRF_A2 s_A1 allTags. - - Definition ESPADA_TSet_8_2 := - [allQs, s_A1] <-$2 ESPADA_TSet_PRF_A1; - [allTags, _] <-$2 oracleMap _ _ (F_bar_random lambda) nil allQs; - ESPADA_TSet_PRF_A2 s_A1 allTags. - - - Theorem ESPADA_TSet_8_1_equiv : - Pr[gInst ESPADA_TSet_7] == Pr[ESPADA_TSet_8_1]. - - unfold ESPADA_TSet_7, ESPADA_TSet_8_1. - unfold ESPADA_TSet_PRF_A1, ESPADA_TSet_PRF_A2. - inline_first. - comp_skip. - inline_first. - comp_simp. - inline_first. - comp_skip. - comp_simp. - comp_skip. - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_eq. - intuition. - subst. - eapply comp_spec_eq_refl. - - eapply eqRat_refl. - - Qed. - - Theorem ESPADA_TSet_8_2_equiv : - | Pr[ESPADA_TSet_8_1] - Pr[ESPADA_TSet_8_2] | <= - (PRF_NA_Advantage (Rnd lambda) (Rnd lambda) F_bar _ _ ESPADA_TSet_PRF_A1 ESPADA_TSet_PRF_A2 ). - - eapply leRat_refl. - Qed. - - Theorem ESPADA_TSet_8_equiv : - Pr[ESPADA_TSet_8_2] == Pr[gInst1 ESPADA_TSet_8]. - - unfold ESPADA_TSet_8_2, ESPADA_TSet_8. - unfold ESPADA_TSet_PRF_A1, ESPADA_TSet_PRF_A2. - inline_first. - do 2 (comp_skip; comp_simp). - eapply eqRat_refl. - - Qed. - - Definition ESPADA_TSet_9_1 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - allTags <-$ ([allTags, _] <-$2 oracleMap _ _ (F_bar_random lambda) nil allQs; ret allTags); - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun x => nth x allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_9_1_equiv : - Pr[gInst1 ESPADA_TSet_8] == Pr[ESPADA_TSet_9_1]. - - unfold ESPADA_TSet_8, ESPADA_TSet_9_1. - - comp_skip; comp_simp. - inline_first. - comp_skip. - comp_simp. - reflexivity. - - Qed. - - Theorem ESPADA_TSet_9_equiv : - Pr[ESPADA_TSet_9_1] == Pr[gInst1 ESPADA_TSet_9]. - - unfold ESPADA_TSet_9_1, ESPADA_TSet_9. - - comp_skip; comp_simp. - comp_skip. - unfold oracleMap. - - symmetry. - rewrite <- evalDist_right_ident. - unfold F_bar_random. - unfold randomFunc. - - eapply eqRat_trans. - comp_skip; comp_simp. - eapply comp_spec_eq_impl_eq. - eapply compMap_fold_equiv. - eapply eqRat_refl. - - unfold compMap_fold. - - eapply comp_spec_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply (@compFold_spec _ _ _ (fun lsa p1 p2 => [rs, s] <-2 p2; p1 = rs /\ NoDup lsa /\ (forall a, In a lsa -> (s # a) = None))); intuition. - - eapply combineKeywords_NoDup. - eapply removeDups_NoDup. - unfold toW. - eapply removeDups_NoDup. - destruct d. - intuition. - subst. - - inversion H0; clear H0; subst. - rewrite H3. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_ret; intuition. - case_eq (eqb a1 a0); intuition. - rewrite eqb_leibniz in H7. - subst. - intuition. - simpl. - rewrite H7. - eapply H3. - simpl. - intuition. - simpl; intuition. - - intuition. - destruct b. - intuition; subst. - eapply comp_spec_ret; intuition. - - reflexivity. - - Qed. - - Definition ESPADA_TSet_10_1 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - allTags <-$ ( - q_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - other_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (skipn (length (removeDups _ q)) allQs); - ret (q_tags ++ other_tags)); - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun x => nth x allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Definition ESPADA_TSet_10_2 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - q_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - other_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (skipn (length (removeDups _ q)) allQs); - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip (q_tags ++ other_tags) ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun x => nth x q_tags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Definition ESPADA_TSet_10_3 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - q_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - other_tags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (skipn (length (removeDups _ q)) allQs); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip q_tags q_ts); - other_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip other_tags other_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun x => nth x q_tags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_10_1_equiv : - Pr[gInst1 ESPADA_TSet_9] == Pr[ESPADA_TSet_10_1]. - - unfold ESPADA_TSet_9, ESPADA_TSet_10_1. - - do 2 (comp_skip; comp_simp). - unfold combineKeywords. - rewrite skipn_app. - eapply compMap_app. - - reflexivity. - Qed. - - Theorem ESPADA_TSet_10_2_equiv : - Pr[ESPADA_TSet_10_1] == Pr[ESPADA_TSet_10_2]. - - unfold ESPADA_TSet_10_1, ESPADA_TSet_10_2. - - comp_skip; comp_simp. - inline_first. - comp_skip; comp_simp. - inline_first. - comp_skip. - intuition. - - comp_skip. - eapply eqRat_eq. - f_equal. - f_equal. - f_equal. - repeat rewrite map_map. - eapply map_ext_in. - intuition. - - erewrite app_nth1. - trivial. - eapply compMap_length in H0. - rewrite H0. - unfold combineKeywords. - rewrite lookupIndex_app. - eapply lookupIndex_lt_length. - eapply in_removeDups. - trivial. - eapply in_removeDups. - trivial. - - Qed. - - - Theorem ESPADA_TSet_10_equiv : - Pr[ESPADA_TSet_10_2] == Pr[gInst1 ESPADA_TSet_10]. - - unfold ESPADA_TSet_10_2, ESPADA_TSet_10. - comp_skip; comp_simp. - comp_skip; comp_simp. - comp_swap_r. - comp_skip; comp_simp. - - unfold combineKeywords. - rewrite map_app. - rewrite <- zip_app. - rewrite map_app. - - rewrite flatten_app. - - eapply eqRat_trans. - comp_skip. - eapply compFold_app. - eapply eqRat_refl. - inline_first. - comp_skip. - - eapply eqRat_refl. - comp_skip. - rewrite skipn_app. - eapply eqRat_refl. - reflexivity. - - rewrite map_length. - eapply compMap_length; eauto. - rewrite map_length. - eapply eq_trans. - eapply compMap_length. - eauto. - unfold combineKeywords. - rewrite skipn_app. - trivial. - Qed. - - Definition ESPADA_TSet_11_1 := - [t, q, s_A] <-$3 A1; - - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - otherTags_ts <-$ compMap _ (fun ls => tag <-$ {0, 1} ^ lambda; ret (tag, ls)) other_ts; - other_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) otherTags_ts; - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes (flatten other_ts_recsList); - - tSet <- getTSet_opt loopRes; - - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - - Lemma ESPADA_TSet_11_1_equiv : - Pr[gInst1 ESPADA_TSet_10] == Pr[ESPADA_TSet_11_1]. - - unfold ESPADA_TSet_10, ESPADA_TSet_11_1. - unfold initFree. - - do 3 (comp_skip; comp_simp). - - eapply comp_spec_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply (@compMap_spec _ _ _ _ _ _ _ (fun t p => t = fst p)); intuition. - eapply list_pred_map_r. - eapply list_pred_eq. - - intuition. - destruct H4. - intuition; subst. - - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - comp_skip; try eapply (oneVector lambda). - eapply comp_spec_ret; intuition. - - intuition. - eapply comp_spec_seq; eauto with inhabited. - eapply compFold_eq. - - eapply list_pred_eq_gen. - f_equal. - f_equal. - - eapply compMap_support in H3. - - eapply list_pred_eq_impl_eq . - eapply list_pred_impl. - eapply list_pred_zip_l; eauto. - - eapply compMap_support in H2. - eapply list_pred_map_r. - - eapply list_pred_symm. - eapply H2. - - intuition. - assert ((fun b a => True) b0 a1). - trivial. - eapply H7. - - intuition. - simpl in *. - subst. - eapply H9. - - intuition. - - repeat simp_in_support. - trivial. - - intuition. - subst. - eapply comp_spec_eq_refl. - - intuition. - subst. - eapply comp_spec_consequence. - eapply comp_spec_eq_refl. - intuition; subst; intuition. - Qed. - - - Definition ESPADA_TSet_11_2 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - ins_F <- map (fun ls => allNatsLt (length ls)) other_ts; - outs_F <-$ ( - otherTags_ins <-$ compMap _ (fun ls => tag <-$ {0, 1} ^ lambda; ret (tag, ls)) ins_F; - ret map (fun tag_ts => [tag,ins] <-2 tag_ts; map (F tag) ins) otherTags_ins); - other_ts_recsList <- map (fun p => [outs_F_list, ls] <-2 p; numberedMap (fun i len p => [out_F, s_i] <-2 p; [b, bigL, bigK] <-3 out_F; (i, len, s_i, b, bigL, bigK)) (zip outs_F_list ls)) (zip outs_F other_ts); - - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - Theorem ESPADA_TSet_11_2_equiv : - Pr[ESPADA_TSet_11_1] == Pr[ESPADA_TSet_11_2]. - - unfold ESPADA_TSet_11_1, ESPADA_TSet_11_2. - - do 3 (comp_skip; comp_simp). - - inline_first. - - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply (@compMap_spec _ _ _ _ _ _ _ (fun p1 p2 => fst p1 = fst p2 /\ snd p2 = (allNatsLt (length (snd p1))))); intuition. - eapply list_pred_map_r. - eapply list_pred_eq. - - comp_skip; try eapply (oneVector lambda). - destruct H4; intuition; subst. - eapply comp_spec_ret; intuition. - - intuition. - comp_simp. - eapply comp_spec_seq; eauto with inhabited. - eapply compFold_eq. - - eapply list_pred_eq_gen. - - eapply flatten_eq. - eapply list_pred_eq_gen. - - eapply map_eq. - - eapply list_pred_impl. - eapply list_pred_zip_r. - eapply list_pred_map_l. - - eapply (@compMap_support _ _ (fun ls p => ls = snd p)) in H3. - - eapply list_pred_symm in H3. - eapply list_pred_map_r_if in H3. - eapply H3. - - intuition. - repeat simp_in_support. - trivial. - - eapply list_pred_map_l. - eapply list_pred_symm in H4. - eauto. - - eapply (compMap_support (fun a p => a = snd p)) in H2. - eapply H2. - - intuition. - repeat simp_in_support. - trivial. - - intuition. - - Ltac exdest := - match goal with - | [H : exists a : _, _ |- _ ] => destruct H - end. - - repeat (try exdest; intuition). - destruct x1. - destruct x2. - destruct b1; simpl in *; subst. - - eapply map_eq. - eapply list_pred_impl. - - eapply list_pred_zip_r. - - eapply list_pred_allNatsLt. - assert (length (zip (map (F b2) (allNatsLt (length b0))) b0) = length (zip (allNatsLt (length b0)) b0)). - repeat rewrite zip_length. - rewrite map_length. - trivial. - eapply allNatsLt_length. - rewrite map_length. - eapply allNatsLt_length. - rewrite H6. - eapply list_pred_allNatsLt. - - eapply list_pred_zip_l. - eapply list_pred_map_l. - eapply list_pred_allNatsLt. - eapply list_pred_map_l. - eapply list_pred_zip_r. - eapply list_pred_allNatsLt. - eapply list_pred_eq. - eapply list_pred_symm. - eapply list_pred_allNatsLt. - eapply list_pred_zip_r. - eapply list_pred_allNatsLt. - eapply list_pred_allNatsLt. - eapply list_pred_eq. - - intuition. - repeat (try exdest; intuition). - destruct b5; simpl in *; subst. - destruct p; simpl in *; subst. - - remember (F b2 x1) as v1. - destruct v1. - destruct p. - repeat f_equal. - - specialize (H6 (n, b6)). - - apply nth_zip in H6. - intuition. - subst. - - eapply nth_allNatsLt. - - eapply allNatsLt_length. - - rewrite zip_length. - rewrite map_length. - rewrite allNatsLt_length. - trivial. - - rewrite map_length. - eapply allNatsLt_length. - - intuition. - subst. - eapply comp_spec_eq_refl. - intuition. - subst. - eapply comp_spec_eq_refl. - Qed. - - Theorem ESPADA_TSet_11_equiv : - Pr[ESPADA_TSet_11_2] == Pr[gInst1 ESPADA_TSet_11]. - - unfold ESPADA_TSet_11_2, ESPADA_TSet_11. - unfold initFree. - do 4 (comp_skip; comp_simp). - - specialize (@compMap_map_fission_eq (list nat) (Bvector lambda * list nat) (list (nat * Bvector lambda * Bvector (S lambda))) (list (nat * Bvector lambda * Bvector (S lambda))) _ _ _ - - (map (fun ls : list (Bvector lambda) => allNatsLt (length ls)) - (map (arrayLookupList (list_EqDec bool_EqDec) t) - (skipn (length (removeDups (list_EqDec bool_EqDec) l)) - (combineKeywords (removeDups (list_EqDec bool_EqDec) l) - (toW t))))) - (fun ls => tag <-$ {0,1 } ^ lambda; ret (tag, ls)) - (fun ls : list nat => tag <-$ { 0 , 1 }^lambda; ret map (F tag) ls) - (fun tag_ts : Bvector lambda * list nat => - [tag, ins]<-2 tag_ts; map (F tag) ins) - (fun x => x) - ); - intuition. - - eapply comp_spec_eq_impl_eq. - eapply comp_spec_eq_trans. - eapply H2. - intuition. - comp_skip; try eapply (oneVector lambda). - eapply comp_spec_ret; intuition. - - eapply comp_spec_consequence. - eapply compMap_spec. - eapply list_pred_eq. - intuition. - subst. - inline_first. - comp_skip. - - eapply comp_spec_eq_refl. - intuition. - eapply list_pred_eq_impl_eq; trivial. - reflexivity. - - Qed. - - Definition ESPADA_TSet_IPRF_A1 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - ins_F <- map (fun ls => allNatsLt (length ls)) other_ts; - ret (ins_F, (s_A, qTags, loopRes, other_ts, allQs, q)). - - Definition ESPADA_TSet_IPRF_A2 s outs_F := - let '(s_A, qTags, loopRes, other_ts, allQs, q) := s in - other_ts_recsList <- map (fun p => [locs, ls] <-2 p; numberedMap (fun i len loc_s_i => [loc, s_i] <-2 loc_s_i; [b, bigL, bigK] <-3 loc; (i, len, s_i, b, bigL, bigK)) (zip locs ls)) (zip outs_F other_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Definition ESPADA_TSet_12_1 := - [ins_F, s] <-$2 ESPADA_TSet_IPRF_A1; - outs_F <-$ compMap _ (fun ls => k <-$ {0, 1}^lambda; ret (map (F k) ls)) ins_F; - ESPADA_TSet_IPRF_A2 s outs_F. - - Definition ESPADA_TSet_12_2 := - [ins_F, s] <-$2 ESPADA_TSet_IPRF_A1; - outs_F <-$ compMap _ (fun lsD => [lsR, _] <-$2 oracleMap _ _ (F_random lambda bigB) nil lsD; ret lsR) ins_F; - ESPADA_TSet_IPRF_A2 s outs_F. - - Theorem ESPADA_TSet_12_1_equiv : - Pr[gInst1 ESPADA_TSet_11] == Pr[ESPADA_TSet_12_1]. - - unfold ESPADA_TSet_11, ESPADA_TSet_12_1. - unfold ESPADA_TSet_IPRF_A1, ESPADA_TSet_IPRF_A2. - - inline_first; comp_skip. - destruct x. - destruct p. - do 2 (inline_first; comp_skip). - reflexivity. - comp_ret_r. - comp_skip. - reflexivity. - - Qed. - - Theorem ESPADA_TSet_12_2_equiv : - | Pr[ESPADA_TSet_12_1] - Pr[ESPADA_TSet_12_2] | <= - (PRF_NAI_Advantage (Rnd lambda) (RndF_range lambda bigB) F _ _ ESPADA_TSet_IPRF_A1 ESPADA_TSet_IPRF_A2 ). - - reflexivity. - - Qed. - - Theorem ESPADA_TSet_12_equiv : - Pr[ESPADA_TSet_12_2] == Pr[gInst1 ESPADA_TSet_12]. - - unfold ESPADA_TSet_12, ESPADA_TSet_12_2. - unfold ESPADA_TSet_IPRF_A1, ESPADA_TSet_IPRF_A2. - - repeat (try inline_first; try comp_skip). - (* comp_simp is slow again *) - destruct x. - destruct p. - repeat (try inline_first; try comp_skip). - eapply eqRat_refl. - reflexivity. - reflexivity. - Qed. - - Theorem ESPADA_TSet_13_equiv : - Pr[gInst1 ESPADA_TSet_12] == Pr[gInst1 ESPADA_TSet_13]. - - unfold ESPADA_TSet_12, ESPADA_TSet_13. - unfold initFree. - - do 4 (comp_skip; comp_simp). - - eapply compMap_eq. - eapply list_pred_map_l. - eapply list_pred_eq. - - intuition. - destruct H2. - intuition; subst. - unfold oracleMap. - symmetry. - eapply comp_spec_eq_impl_eq. - eapply comp_spec_eq_trans. - eapply compMap_fold_equiv. - unfold compMap_fold. - eapply comp_spec_eq_trans. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - eapply comp_spec_eq_symm. - - eapply comp_spec_seq; try eapply nil. - eapply (@compFold_spec' _ _ _ _ -(fun lsa lsb p1 p2 => [rs, s] <-2 p1; rs = p2 /\ NoDup lsa /\ forall a, In a lsa -> (s # a) = None)); intuition. - - eapply allNatsLt_length. - eapply allNatsLt_NoDup. - - subst. - inversion H3; clear H3; subst. - unfold F_random. - unfold randomFunc. - rewrite H6. - unfold RndF_range. - inline_first. - comp_skip. - inline_first. - do 2 comp_skip. - eapply comp_spec_ret; intuition. - simpl. - case_eq (eqb a3 a0); intuition. - rewrite eqb_leibniz in H13. - subst. - intuition. - simpl; intuition. - - intuition. - subst. - eapply comp_spec_eq_refl. - - Qed. - - - Definition ESPADA_TSet_14_1 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (fun q => ts <- arrayLookupList _ t q; numberedMap (fun i len s_i => (i, len, s_i)) ts) (skipn (length (removeDups _ q)) allQs); - outs_F <-$ compMap _ (fun ls => compMap _ (fun _ => RndF_range lambda bigB) ls) other_ts; - other_ts_recsList <- map (fun p => [locs, ls] <-2 p; map (fun loc_p => [loc, p] <-2 loc_p; [i, len, s_i] <-3 p; [b, bigL, bigK] <-3 loc; (i, len, s_i, b, bigL, bigK)) (zip locs ls)) (zip outs_F other_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - Theorem ESPADA_TSet_14_1_equiv : - Pr[gInst1 ESPADA_TSet_13] == Pr[ESPADA_TSet_14_1]. - - unfold ESPADA_TSet_13, ESPADA_TSet_14_1. - - do 3 (comp_skip; comp_simp). - reflexivity. - - assert ( forall x1, - eqRat - (evalDist - (compMap - (list_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda)))) - (fun ls : list (Bvector lambda) => - compMap - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) - (fun _ : Bvector lambda => RndF_range lambda bigB) ls) - (map (arrayLookupList (list_EqDec bool_EqDec) t) - (skipn (length (removeDups (list_EqDec bool_EqDec) l)) - (combineKeywords (removeDups (list_EqDec bool_EqDec) l) - (toW t))))) x1) - (evalDist - (compMap - (list_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda)))) - (fun ls : list (prod (prod nat nat) (Bvector lambda)) => - compMap - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) - (fun _ : prod (prod nat nat) (Bvector lambda) => - RndF_range lambda bigB) ls) - (map - (fun q : Blist => - numberedMap - (fun (i len : nat) (s_i : Bvector lambda) => - pair (pair i len) s_i) - (arrayLookupList (list_EqDec bool_EqDec) t q)) - (skipn (length (removeDups (list_EqDec bool_EqDec) l)) - (combineKeywords (removeDups (list_EqDec bool_EqDec) l) - (toW t))))) x1) - ). - - eapply (@compMap_eq). - - eapply list_pred_map_r. - eapply list_pred_map_l. - eapply list_pred_eq. - - intuition. - repeat (try exdest; intuition). - subst. - eapply compMap_eq. - unfold numberedMap. - eapply list_pred_map_r. - eapply list_pred_zip_r. - eapply list_pred_allNatsLt. - eapply list_pred_allNatsLt. - eapply list_pred_eq. - - intuition. - - comp_skip. - - comp_skip. - eapply comp_spec_eq_impl_eq. - eapply (@compFold_eq _ _ eq). - eapply list_pred_flatten_both. - eapply list_pred_impl. - eapply list_pred_eq_gen. - eapply map_eq. - - assert (In x1 - (getSupport (compMap - (list_EqDec - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda)))) - (fun ls : list (prod (prod nat nat) (Bvector lambda)) => - compMap - (pair_EqDec (pair_EqDec nat_EqDec (Bvector_EqDec lambda)) - (Bvector_EqDec (S lambda))) - (fun _ : prod (prod nat nat) (Bvector lambda) => - RndF_range lambda bigB) ls) - (map - (fun q : Blist => - numberedMap - (fun (i len : nat) (s_i : Bvector lambda) => - pair (pair i len) s_i) - (arrayLookupList (list_EqDec bool_EqDec) t q)) - (skipn (length (removeDups (list_EqDec bool_EqDec) l)) - (combineKeywords (removeDups (list_EqDec bool_EqDec) l) - (toW t)))))) - ). - - eapply getSupport_In_evalDist. - intuition. - eapply getSupport_In_evalDist. - eauto. - rewrite H2. - intuition. - - apply (@compMap_support _ _ (fun p1 p2 => length p1 = length p2)) in H3. - apply (@compMap_support _ _ (fun p1 p2 => length p1 = length p2)) in H4. - - - eapply list_pred_impl. - eapply list_pred_zip_l. - - eapply list_pred_symm. - eapply H3. - - eapply list_pred_zip_r. - eapply list_pred_symm. - - eapply H4. - eapply list_pred_eq. - - eapply H4. - eapply list_pred_zip_r. - eapply list_pred_symm. - eapply H4. - eapply list_pred_symm. - eapply H3. - - eapply list_pred_map_l. - eapply list_pred_map_r. - eapply list_pred_eq. - - intuition. - repeat (try exdest; intuition). - destruct b0. - simpl in *; subst. - - unfold numberedMap. - eapply map_eq. - - eapply list_pred_impl. - eapply list_pred_zip_l. - eapply list_pred_allNatsLt. - eapply list_pred_zip_r. - eapply list_pred_map_r. - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - - eapply list_pred_zip_r. - eapply list_pred_allNatsLt. - - rewrite H11. - eapply list_pred_allNatsLt. - eapply list_pred_I. - eauto. - eapply allNatsLt_length. - rewrite zip_length. - eapply list_pred_symm. - eapply list_pred_allNatsLt. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - auto. - eapply allNatsLt_length. - - eapply list_pred_map_l. - eapply list_pred_zip_l. - eapply list_pred_allNatsLt. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - - rewrite zip_length. - rewrite H11. - eapply list_pred_eq. - rewrite H11. - trivial. - eapply allNatsLt_length. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - rewrite zip_length. - rewrite <- H11. - eapply list_pred_symm. - eapply list_pred_allNatsLt. - auto. - eapply allNatsLt_length. - - eapply list_pred_zip_r. - eapply list_pred_map_r. - eapply list_pred_zip_r. - eapply list_pred_allNatsLt. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - rewrite H11. - eapply list_pred_allNatsLt. - - eapply allNatsLt_length. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - - eapply allNatsLt_length. - eapply list_pred_zip_r. - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - eapply allNatsLt_length. - eapply list_pred_eq. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - eapply allNatsLt_length. - - eapply list_pred_map_l. - eapply list_pred_zip_l. - eapply list_pred_allNatsLt. - eapply list_pred_zip_r. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - eapply allNatsLt_length. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - rewrite H11. - eapply list_pred_symm. - eapply list_pred_allNatsLt. - eapply allNatsLt_length. - eapply list_pred_symm. - eapply list_pred_allNatsLt. - - eapply list_pred_zip_r. - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - eapply allNatsLt_length. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - eapply list_pred_I. - auto. - eapply allNatsLt_length. - - eapply list_pred_eq. - - intuition. - repeat (try exdest; intuition). - destruct b3. - destruct x3. - destruct x5. - destruct x6. - destruct x7. - simpl in *; subst. - pairInv. - - repeat f_equal. - - unfold numberedMap in H11. - rewrite map_length in H11. - rewrite zip_length in H11. - rewrite allNatsLt_length in H11. - rewrite zip_length. - rewrite H11. - trivial. - rewrite H11. - trivial. - eapply allNatsLt_length. - - intuition. - apply compMap_length in H6. - auto. - - intuition. - apply compMap_length in H6. - auto. - - intuition. - - subst. - eapply list_pred_eq. - - intuition; subst. - eapply comp_spec_eq_refl. - reflexivity. - Qed. - - (* Now I can flatten the map *) - Definition ESPADA_TSet_14_2 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (fun q => ts <- arrayLookupList _ t q; numberedMap (fun i len s_i => (i, len, s_i)) ts) (skipn (length (removeDups _ q)) allQs); - outs_F <-$ compMap _ (fun _ => RndF_range lambda bigB) (flatten other_ts); - other_ts_recsList <- map (fun e => [p1, p2] <-2 e; [i, len, s_i] <-3 p1; [b, bigL, bigK] <-3 p2; (i, len, s_i, b, bigL, bigK)) (zip (flatten other_ts) outs_F); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes other_ts_recsList; - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_14_2_equiv : - Pr[ESPADA_TSet_14_1] == Pr[ESPADA_TSet_14_2]. - - unfold ESPADA_TSet_14_1, ESPADA_TSet_14_2. - - do 3 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply compMap_flatten; intuition. - intuition. - subst. - - eapply comp_spec_seq; eauto with inhabited. - eapply compFold_eq. - eapply list_pred_eq_gen. - - rewrite flatten_map_pair_eq. - - repeat rewrite unzip_zip_inv. - unfold fst. - unfold snd. - - eapply map_eq. - - eapply list_pred_impl. - eapply list_pred_zip_eq_rev. - - eapply length_flatten_eq. - eapply list_pred_symm. - eapply compMap_support. - eauto. - - intuition. - apply compMap_length in H5. - trivial. - - intuition. - destruct b4. - destruct p0. - destruct p. - intuition. - repeat pairInv. - trivial. - - apply compMap_length in H2. - eapply H2. - - intuition. - apply (@compMap_support _ _ (fun ls1 ls2 => length ls1 = length ls2)) in H2. - - eapply list_pred_zip_in in H2. - eauto. - - eapply in_zip_swap;intuition. - symmetry. - eapply list_pred_length_eq. - eauto. - - intuition. - apply compMap_length in H6. - auto. - - intuition. - subst. - eapply comp_spec_eq_refl. - - intuition. - subst. - eapply comp_spec_eq_refl. - Qed. - - - Definition ESPADA_TSet_14_3 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (fun q => ts <- arrayLookupList _ t q; numberedMap (fun i len s_i => (i, len, s_i)) ts) (skipn (length (removeDups _ q)) allQs); - outs_F <-$ compMap _ (fun p => [i, len, s_i] <-3 p; [b, bigL, bigK] <-$3 RndF_range lambda bigB; ret (i, len, s_i, b, bigL, bigK)) (flatten other_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) loopRes outs_F; - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - Theorem ESPADA_TSet_14_3_equiv : - Pr[ESPADA_TSet_14_2] == Pr[ESPADA_TSet_14_3]. - - unfold ESPADA_TSet_14_2, ESPADA_TSet_14_3. - - do 3 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - eapply comp_spec_seq; eauto with inhabited. - eapply (@compMap_spec _ _ _ _ _ _ _ (fun p1 p2 => [b1, bigL1, bigK1] <-3 p1; let '(i, len, s_i, b2, bigL2, bigK2) := p2 in b1 = b2 /\ bigL1 = bigL2 /\ bigK1 = bigK2)); intuition. - - eapply list_pred_eq. - - subst. - eapply comp_spec_eq_trans_l. - eapply comp_spec_eq_symm. - eapply comp_spec_right_ident. - - comp_skip; try eapply (oneVector (S lambda)). - eapply comp_spec_ret; intuition. - - intuition. - eapply (@compMap_support _ _ (fun p1 p2 => [i1, len1, s_i1] <-3 p1; let '(i2, len2, s_i2, b, bigL, bigK) := p2 in i1 = i2 /\ len1 = len2 /\ s_i1 = s_i2 )) in H3. - eapply (@compMap_support _ _ (fun p1 p2 => True)) in H2. - - eapply comp_spec_seq; eauto with inhabited. - eapply (@compFold_eq). - - eapply list_pred_map_l. - eapply list_pred_zip_l. - eapply list_pred_I. - eapply eq_trans. - eapply list_pred_length_eq. - eapply H2. - trivial. - eapply H3. - eauto. - - intuition. - - repeat (try exdest; intuition). - destruct x1. - destruct a3. - simpl in *. - destruct p0. - destruct p1. - destruct p0. - destruct p1. - destruct p0. - destruct p0. - intuition; subst. - destruct p. - destruct p. - pairInv. - simpl in *. - - intuition; subst. - eapply comp_spec_eq_refl. - - intuition. - subst. - eapply comp_spec_eq_refl. - - intuition. - intuition. - - repeat simp_in_support. - destruct x1. - destruct p. - repeat simp_in_support. - intuition. - - Qed. - - Definition ESPADA_TSet_14_4 := - [t, q, s_A] <-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_6 lambda)) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (fun q => ts <- arrayLookupList _ t q; numberedMap (fun i len s_i => (i, len, s_i)) ts) (skipn (length (removeDups _ q)) allQs); - loopRes <-$ compFold _ (foldBody_option _ (@ESPADA_TSetSetup_tLoop_14 lambda bigB)) loopRes (flatten other_ts); - tSet <- getTSet_opt loopRes; - inds <- - (foreach (x in q) - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) allQs x 0); - tags <- map (fun x => nth x qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - Theorem ESPADA_TSet_14_4_equiv : - Pr[ESPADA_TSet_14_3] == Pr[ESPADA_TSet_14_4]. - - unfold ESPADA_TSet_14_3, ESPADA_TSet_14_4. - - do 3 (comp_skip; comp_simp). - rewrite <- evalDist_assoc. - comp_skip. - symmetry. - eapply comp_spec_eq_impl_eq. - eapply fold_map_fission_spec; - intuition. - - subst. - unfold foldBody_option. - unfold ESPADA_TSetSetup_tLoop_6, ESPADA_TSetSetup_tLoop_14. - - destruct e. - comp_simp. - unfold RndF_range. - inline_first. - comp_skip. - inline_first. - comp_skip. - inline_first. - comp_skip. - comp_simp. - eapply comp_spec_eq_refl. - - comp_irr_r. - eapply pair_EqDec; intuition. - unfold RndF_range. - wftac. - - intuition. - - Qed. - - Theorem ESPADA_TSet_14_equiv : - Pr[ESPADA_TSet_14_4] == Pr[gInst1 ESPADA_TSet_14]. - - unfold ESPADA_TSet_14_4, ESPADA_TSet_14. - - do 3 (comp_skip; comp_simp). - - eapply comp_spec_eq_impl_eq. - eapply (@compFold_eq _ _ - (fun p1 p2 => let '(i1, len1, v1, b, bigL, bigK) := p1 in - let '(i2, len2, tag, v2) := p2 in - i1 = i2 /\ len1 = len2 /\ v1 = v2 /\ (b, bigL, bigK) = F tag i1)). - eapply list_pred_flatten_both. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - - intuition. - subst. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - eapply list_pred_impl. - eapply list_pred_eq. - - intuition. - subst. - remember (F a1 a2) as v1. - destruct v1. - destruct p. - intuition. - - intuition. - destruct a2. - destruct p. - destruct p. - intuition. - subst. - unfold ESPADA_TSetSetup_tLoop_6, ESPADA_TSetSetup_tLoop_4. - - unfold foldBody_option. - destruct acc. - destruct p. - remember (F b5 n) as v1. - destruct v1. - destruct p. - - comp_simp. - pairInv. - eapply comp_spec_eq_refl. - eapply comp_spec_eq_refl. - - comp_skip. - eapply comp_spec_eq_impl_eq. - eapply compFold_eq. - eapply list_pred_flatten_both. - eapply list_pred_map_r'. - eapply list_pred_map_l'. - eapply list_pred_map_r'. - eapply list_pred_impl. - eapply list_pred_eq. - - intuition. - subst. - eapply list_pred_eq. - - intuition. - subst. - eapply comp_spec_eq_refl. - reflexivity. - Qed. - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - (forall db q s_A, - In (db, q, s_A) (getSupport A1) -> - length (combineKeywords (removeDups _ q) (toW db)) <= maxKeywords)%nat. - - Variable PRF_NA_Adv : Rat. - Hypothesis PRF_NA_Adv_correct_1 : - (PRF_NA_Advantage (Rnd lambda) (Rnd lambda) F_bar _ _ ESPADA_TSet_PRF_A1 ESPADA_TSet_PRF_A2 ) <= PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_2: - forall i, - PRF_NA_Advantage (Rnd lambda) (@RndF_range lambda bigB) F _ _ - (B1 nil _ _ ESPADA_TSet_IPRF_A1 i) - (B2 (fun lsD => k <-$ {0, 1}^lambda; ret (map (F k) lsD)) - (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func (@RndF_range lambda bigB) _) nil lsD; ret lsR) - _ ESPADA_TSet_IPRF_A2 - ) - <= PRF_NA_Adv. - - Theorem ESPADA_TSet_once_secure : - | Pr[TSetReal _ (@ESPADA_TSetSetup_once lambda bigB bigS F F_bar) (@ESPADA_TSetGetTag lambda F_bar) A1 A2] - - Pr[TSetIdeal (@L_T lambda) A1 A2 (@ESPADA_TSet_Sim_once lambda bigB bigS F)] | - <= PRF_NA_Adv + (maxKeywords / 1) * PRF_NA_Adv. - - rewrite ESPADA_TSet_1_equiv. - rewrite ESPADA_TSet_2_equiv. - rewrite ESPADA_TSet_3_1_equiv. - rewrite ESPADA_TSet_3_2_equiv. - rewrite ESPADA_TSet_3_equiv. - rewrite ESPADA_TSet_4_equiv. - rewrite ESPADA_TSet_6_equiv. - rewrite ESPADA_TSet_7_equiv. - rewrite ESPADA_TSet_8_1_equiv. - eapply leRat_trans. - eapply ratDistance_le_trans. - eapply ESPADA_TSet_8_2_equiv. - rewrite ESPADA_TSet_8_equiv. - rewrite ESPADA_TSet_9_1_equiv. - rewrite ESPADA_TSet_9_equiv. - rewrite ESPADA_TSet_10_1_equiv. - rewrite ESPADA_TSet_10_2_equiv. - rewrite ESPADA_TSet_10_equiv. - rewrite ESPADA_TSet_11_1_equiv. - rewrite ESPADA_TSet_11_2_equiv. - rewrite ESPADA_TSet_11_equiv. - rewrite ESPADA_TSet_12_1_equiv. - eapply ratDistance_le_trans. - eapply ESPADA_TSet_12_2_equiv. - eapply eqRat_impl_leRat. - rewrite <- ratIdentityIndiscernables. - rewrite ESPADA_TSet_12_equiv. - rewrite ESPADA_TSet_13_equiv. - rewrite ESPADA_TSet_14_1_equiv. - rewrite ESPADA_TSet_14_2_equiv. - rewrite ESPADA_TSet_14_3_equiv. - rewrite ESPADA_TSet_14_4_equiv. - rewrite ESPADA_TSet_14_equiv. - rewrite ESPADA_TSet_15_equiv. - rewrite ESPADA_TSet_16_equiv. - rewrite ESPADA_TSet_17_equiv. - rewrite ESPADA_TSet_18_equiv. - rewrite ESPADA_TSet_19_equiv. - rewrite Ideal_equiv. - - intuition. - - rewrite <- ratAdd_0_r. - eapply ratAdd_leRat_compat; intuition. - rewrite PRF_NA_impl_NAI. - reflexivity. - intuition. - unfold ESPADA_TSet_IPRF_A1 in *. - repeat simp_in_support. - destruct x. - destruct p. - repeat simp_in_support. - repeat rewrite map_length. - - Theorem skipn_length : - forall (A : Type)(ls : list A)(n : nat), - (length (skipn n ls) <= length ls - n)%nat. - - induction ls; intuition; simpl in *. - destruct n; simpl; trivial. - destruct n; simpl; trivial. - - Qed. - - rewrite skipn_length. - eapply minus_le. - eapply maxKeywords_correct. - eauto. - unfold RndF_range. - wftac. - wftac. - - eauto. - Qed. - - Print Assumptions ESPADA_TSet_once_secure. - - (* State the previous theorem in terms of DistMult so it can be more easily used in the next stage of the proof. *) - - Require Import RepeatCore. - - Definition Repeat_Real t_q := - [t, q] <-2 t_q; - [tSet, k_T] <-$2 ESPADA_TSetSetup_once bigB bigS F F_bar t; - tags <-$ compMap (Bvector_EqDec lambda) (ESPADA_TSetGetTag lambda F_bar k_T) q; - ret (tSet, tags). - - Definition Repeat_Ideal t_q := - [t, q] <-2 t_q; - T_qs <- map (arrayLookupList (list_EqDec bool_EqDec) t) q; - ESPADA_TSet_Sim_once lambda bigB bigS F (@L_T lambda t q) T_qs. - - Theorem ESPADA_TSet_once_eq_DistSingle : - (| Pr[TSetReal _ (@ESPADA_TSetSetup_once lambda bigB bigS F F_bar) (@ESPADA_TSetGetTag lambda F_bar) A1 A2] - - Pr[TSetIdeal (@L_T lambda) A1 A2 (@ESPADA_TSet_Sim_once lambda bigB bigS F)] |) == - DistSingle_Adv Repeat_Real Repeat_Ideal A1 A2. - - unfold DistSingle_Adv, DistSingle_G, TSetReal, TSetIdeal, Repeat_Real, Repeat_Ideal. - - eapply ratDistance_eqRat_compat; - repeat (inline_first; comp_skip; comp_simp; try reflexivity). - Qed. - - - Theorem ESPADA_TSet_DistSingle_secure : - DistSingle_Adv Repeat_Real Repeat_Ideal A1 A2 - <= PRF_NA_Adv + (maxKeywords / 1) * PRF_NA_Adv. - - rewrite <- ESPADA_TSet_once_eq_DistSingle. - apply ESPADA_TSet_once_secure. - - Qed. - -End ESPADA_TSet_once. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Once_Games.v b/src/ESPADA/ESPADA_TSet_Once_Games.v deleted file mode 100644 index f5c1ec4..0000000 --- a/src/ESPADA/ESPADA_TSet_Once_Games.v +++ /dev/null @@ -1,451 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import FCF. -Require Import CompFold. -Require Import TSet. -Require Import RndListElem. -Require Import List. -Require Import RndNat. -Require Import Array. -Require Import PRF. -Require Export ESPADA_TSet. - -Section ESPADA_TSet_Games. - - Variable lambda : posnat. - Variable tSize bigB bigS : posnat. - Variable bucketSize : nat -> (posnat * posnat). - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Definition TSet := (@TSet lambda). - Definition ESPADA_TSetSetup_trial := (@ESPADA_TSetSetup_trial lambda bigB bigS F F_bar). - Definition ESPADA_TSetSetup_tLoop := (@ESPADA_TSetSetup_tLoop lambda F). - Definition ESPADA_TSetSetup_wLoop := (@ESPADA_TSetSetup_wLoop lambda F F_bar). - Definition initFree := (@initFree bigB bigS). - - (* The adversary *) - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable A1 : Comp (T lambda * list Blist * A_State). - Variable A2 : A_State -> (option TSet) * list (Bvector lambda) -> Comp bool. - - Definition getTSet_opt(opt : option (TSet * FreeList)) := - match opt with - | None => None - | Some p => Some (fst p) - end. - - Definition ESPADA_TSetSetup_once t := - [res, k_T] <-$2 ESPADA_TSetSetup_trial t; - ret (getTSet_opt res, k_T). - - Definition randomTSetEntry (acc : TSet * FreeList) := - label <-$ {0, 1} ^ lambda; - value <-$ {0, 1} ^ (S lambda); - [tSet, free] <-2 acc; - (* To match the actual setup routine, we choose b uniformly at random and fail if that bucket is empty. *) - b <-$ [0 .. bigB); - free_b <- nth b free nil; - j <-$? (rndListElem _ free_b); - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - tSet <- tSetUpdate tSet b j (label, value); - ret (tSet, free). - - Definition ESPADA_TSetSetup_Sim_wLoop (tSet_free : TSet * FreeList) e := - [tSet, free] <-2 tSet_free; - [stag, t] <-2 e; - ls <- zip (allNatsLt (length t)) t; - loop_over ((tSet, free), ls) (ESPADA_TSetSetup_tLoop stag (length t)). - - - Definition ESPADA_TSet_Sim_trial (leak : nat * list (nat))(ts : list (list (Bvector lambda))) := - [n, eqPattern ] <-2 leak; - qTags <-$ foreach (_ in (removeDups _ eqPattern)) ({0, 1} ^ lambda); - uniqueTs <- map (fun n => nth (firstIndexOf (EqDec_dec _ ) eqPattern n O) ts nil) (removeDups _ eqPattern); - loopRes <-$ loop_over ((nil, initFree), (zip qTags uniqueTs)) ESPADA_TSetSetup_Sim_wLoop; - loopRes <-$ loop_over* (loopRes, (allNatsLt (minus n (length (flatten uniqueTs))))) (fun acc i => randomTSetEntry acc); - tags <- map (fun i => nth i qTags (Vector.const false lambda)) eqPattern; - ret (loopRes, tags). - - Definition ESPADA_TSet_Sim (leak : nat * list nat)(ts : list (list (Bvector lambda))) := - [trialRes, tags] <-$2 Repeat (ESPADA_TSet_Sim_trial leak ts) (fun p => isSome (fst p)); - ret (getTSet trialRes, tags). - - Definition ESPADA_TSet_Sim_once (leak : nat * list nat)(ts : list (list (Bvector lambda))) := - [trialRes, tags] <-$2 ESPADA_TSet_Sim_trial leak ts; - ret (getTSet_opt trialRes, tags). - - Definition bigN (bigT : T lambda) := - bigW <- toW bigT; - fold_left (fun acc w => acc + length (arrayLookupList _ bigT w))%nat bigW O. - - Definition L_T bigT (q : list Blist) := - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - (bigN bigT, inds). - - (* The sequence of games starts from real and moves toward ideal. *) - - (* Step 1 : Inline and simplify. *) - Definition ESPADA_TSet_1 := - [t, q, s_A]<-$3 A1; - k_T <-$ {0, 1} ^ lambda; - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_wLoop t k_T)) (Some (nil, initFree)) (toW t); - tSet <- getTSet_opt loopRes; - tags <- map (F_bar k_T) q; - A2 s_A (tSet, tags). - - (* Step 2: pull out the stag computation and answer lookup. We switch to the simulator setup loop because it happens to work here. *) - Definition ESPADA_TSet_2 := - [t, q, s_A]<-$3 A1; - k_T <-$ {0, 1} ^ lambda; - stags <- map (F_bar k_T) (toW t); - ts <- map (arrayLookupList _ t) (toW t); - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip stags ts); - tSet <- getTSet_opt loopRes; - tags <- map (F_bar k_T) q; - A2 s_A (tSet, tags). - - (* Step 3: initialize the TSet using all keywords -- including those that are in the the queries but not in the structure. *) - Local Open Scope list_scope. - - - Definition combineKeywords (q1 q2 : list Blist):= - q1 ++ (removePresent (EqDec_dec _) q1 q2). - - Definition ESPADA_TSet_3 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (toW t) (removeDups _ q); - allTags <-$ (k_T <-$ {0, 1} ^ lambda; ret (map (F_bar k_T) allQs)); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - ts <- map (arrayLookupList _ t) allQs; - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip allTags ts); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 4 : flatten the list. *) - Definition numberedMap(A B : Set)(f : nat -> nat -> A -> B)(ls : list A) := - map (fun p => [i, a] <-2 p; f i (length ls) a) (zip (allNatsLt (length ls)) ls). - - Definition ESPADA_TSetSetup_tLoop_4 (acc : TSet * FreeList) (e : nat * nat * Bvector lambda * Bvector lambda) := - [tSet, free] <-2 acc; - let '(i, len, tag, s_i) := e in - [b, bigL, bigK] <-3 F tag i; - free_b <- nth b free nil; - opt_j <-$ rndListElem _ (free_b); - match opt_j with - | None => ret None - | Some j => - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - bet <- if (eq_nat_dec (S i) len) then false else true; - newRecord <- (bigL, BVxor _ (Vector.cons _ bet _ s_i) bigK); - tSet <- tSetUpdate tSet b j newRecord; - ret (Some (tSet, free)) - end. - - Definition ESPADA_TSet_4 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (toW t) (removeDups _ q); - allTags <-$ (k_T <-$ {0, 1} ^ lambda; ret (map (F_bar k_T) allQs)); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - ts <- map (arrayLookupList _ t) allQs; - (* Put all the required information in each answer entry so we can flatten the list. *) - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => (i, len, tag, s_i)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_4) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Todo : fix the numbering *) - - (* Step 6 : factor out the inner PRF*) - Definition ESPADA_TSetSetup_tLoop_6 (acc : TSet * FreeList) e := - [tSet, free] <-2 acc; - let '(i, len, s_i, b, bigL, bigK) := e in - free_b <- nth b free nil; - opt_j <-$ rndListElem _ (free_b); - match opt_j with - | None => ret None - | Some j => - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - bet <- if (eq_nat_dec (S i) len) then false else true; - newRecord <- (bigL, BVxor _ (Vector.cons _ bet _ s_i) bigK); - tSet <- tSetUpdate tSet b j newRecord; - ret (Some (tSet, free)) - end. - - Definition ESPADA_TSet_6 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (toW t) (removeDups _ q); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - allTags <-$ (k_T <-$ {0, 1} ^ lambda; ret (map (F_bar k_T) allQs)); - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 7 : change the initialization order so that queried keywords are initialized first. This makes the order match the simulator. *) - Definition ESPADA_TSet_7 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - allTags <-$ (k_T <-$ {0, 1} ^ lambda; ret (map (F_bar k_T) allQs)); - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - (* Step 8 : replace the first PRF with a random function *) - Definition F_bar_random := (@randomFunc Blist _ (Rnd lambda) _). - Definition ESPADA_TSet_8 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - [allTags, _] <-$2 oracleMap _ _ F_bar_random nil allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 9 : There are no duplicates in the random function inputs, replace the outputs with random values *) - Definition ESPADA_TSet_9 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - allTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) allQs; - ts <- map (arrayLookupList _ t) allQs; - ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip allTags ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten ts_recsList); - tSet <- getTSet_opt loopRes; - tags <- map (fun i => nth i allTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 10 : Split the init loop into two parts. First initialize the queried keywords, then initialize the rest *) - Definition ESPADA_TSet_10 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten q_ts_recsList); - - otherTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (skipn (length (removeDups _ q)) allQs); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - other_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip otherTags other_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) loopRes (flatten other_ts_recsList); - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - tSet <- getTSet_opt loopRes; - A2 s_A (tSet, tags). - - (* Step 11 : We will replace F with a random function only where it is used to initialize non-queried keywords. Start by calling the PRF in a way that matches the definition. *) - Definition ESPADA_TSet_11 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - ins_F <- map (fun ls => allNatsLt (length ls)) other_ts; - outs_F <-$ compMap _ (fun ls => k <-$ {0, 1}^lambda; ret (map (F k) ls)) ins_F; - - other_ts_recsList <- map (fun p => [locs, ls] <-2 p; numberedMap (fun i len loc_s_i => [loc, s_i] <-2 loc_s_i; [b, bigL, bigK] <-3 loc; (i, len, s_i, b, bigL, bigK)) (zip locs ls)) (zip outs_F other_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) loopRes (flatten other_ts_recsList); - - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 12 : Replace PRF F with a random function. *) - Definition RndF_range := - b <-$ [0 .. bigB); - x <-$ {0, 1} ^ lambda; - y <-$ {0, 1} ^ (S lambda); - ret (b, x, y). - Definition F_random := (@randomFunc nat _ (RndF_range) _). - - Definition ESPADA_TSet_12 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten q_ts_recsList); - - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - ins_F <- map (fun ls => allNatsLt (length ls)) other_ts; - outs_F <-$ compMap _ (fun lsD => [lsR, _] <-$2 oracleMap _ _ (F_random) nil lsD; ret lsR) ins_F; - - other_ts_recsList <- map (fun p => [locs, ls] <-2 p; numberedMap (fun i len loc_s_i => [loc, s_i] <-2 loc_s_i; [b, bigL, bigK] <-3 loc; (i, len, s_i, b, bigL, bigK)) (zip locs ls)) (zip outs_F other_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) loopRes (flatten other_ts_recsList); - - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 13 : There are no duplicates in the inputs to the random function (all inputs are the set of natural numbers less than k for some k), so we can replace the outputs with random values. *) - - Definition ESPADA_TSet_13 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => [b, bigL, bigK] <-3 F tag i; (i, len, s_i, b, bigL, bigK)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - outs_F <-$ compMap _ (fun ls => compMap _ (fun _ => RndF_range) ls) other_ts; - - other_ts_recsList <- map (fun p => [locs, ls] <-2 p; numberedMap (fun i len loc_s_i => [loc, s_i] <-2 loc_s_i; [b, bigL, bigK] <-3 loc; (i, len, s_i, b, bigL, bigK)) (zip locs ls)) (zip outs_F other_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_6) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 14 : put the inner PRF calls (and random sampling) back into the inner loops. *) - Definition ESPADA_TSetSetup_tLoop_14 (acc : TSet * FreeList) (e : nat * nat * Bvector lambda) := - [tSet, free] <-2 acc; - let '(i, len, s_i) := e in - b <-$ [0 .. bigB); - bigL <-$ { 0 , 1 }^lambda; - enc_val <-$ ( - bigK <-$ { 0 , 1 }^S lambda; - bet <- if (eq_nat_dec (S i) len) then false else true; - ret (BVxor _ (Vector.cons _ bet _ s_i) bigK)); - free_b <- nth b free nil; - opt_j <-$ rndListElem _ (free_b); - match opt_j with - | None => ret None - | Some j => - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - - newRecord <- (bigL, enc_val); - tSet <- tSetUpdate tSet b j newRecord; - ret (Some (tSet, free)) - end. - - Definition ESPADA_TSet_14 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => (i, len, tag, s_i)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_4) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - other_ts_recsList <- map (fun ls => numberedMap (fun i len e => (i, len, e)) ls) other_ts; - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_14) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - (* Step 15 : Now we can replace the second part with a loop that creates random entries (as in the simulator) *) - Definition ESPADA_TSetSetup_tLoop_15 (acc : TSet * FreeList) (e : nat * nat * Bvector lambda) := - [tSet, free] <-2 acc; - b <-$ [0 .. bigB); - bigL <-$ { 0 , 1 }^lambda; - bigK <-$ { 0 , 1 }^S lambda; - (* let '(len, i, s_i) := e in *) - free_b <- nth b free nil; - opt_j <-$ rndListElem _ (free_b); - match opt_j with - | None => ret None - | Some j => - free <- listReplace free b (removeFirst (EqDec_dec _) free_b j) nil; - tSet <- tSetUpdate tSet b j (bigL, bigK); - ret (Some (tSet, free)) - end. - - Definition ESPADA_TSet_15 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => (i, len, tag, s_i)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_4) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - other_ts_recsList <- map (fun ls => numberedMap (fun i len e => (i, len, e)) ls) other_ts; - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_15) loopRes (flatten other_ts_recsList); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 16 : After the one-time pad argument, many values are no longer used. Simplify. *) - - Definition ESPADA_TSet_16 := - [t, q, s_A]<-$3 A1; - allQs <- combineKeywords (removeDups _ q) (toW t); - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => (i, len, tag, s_i)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_4) (Some (nil, initFree)) (flatten q_ts_recsList); - other_ts <- map (arrayLookupList _ t) (skipn (length (removeDups _ q)) allQs); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry acc)) loopRes (flatten other_ts); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) allQs x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 17 : don't look up answers for the non-queried keywords *) - Definition ESPADA_TSet_17 := - [t, q, s_A]<-$3 A1; - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - q_ts_recsList <- map (fun p => [tag, ls] <-2 p; numberedMap (fun i len s_i => (i, len, tag, s_i)) ls) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ ESPADA_TSetSetup_tLoop_4) (Some (nil, initFree)) (flatten q_ts_recsList); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry acc)) loopRes (allNatsLt (minus (bigN t) (length (flatten q_ts)))); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - - (* Step 18 : un-flatten the list in the first step *) - Definition ESPADA_TSet_18 := - [t, q, s_A]<-$3 A1; - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ q); - q_ts <- map (arrayLookupList _ t) (removeDups _ q); - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry acc)) loopRes (allNatsLt (minus (bigN t) (length (flatten q_ts)))); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* Step 19: use (the body of) the leakage function to produce the required information. *) - Definition ESPADA_TSet_19 := - [t, q, s_A]<-$3 A1; - ts <- map (arrayLookupList _ t) q; - eqPattern <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - qTags <-$ compMap _ (fun _ => {0, 1} ^ lambda) (removeDups _ eqPattern); - - q_ts <- map (fun n => nth (firstIndexOf (EqDec_dec _ ) eqPattern n O) ts nil) (removeDups _ eqPattern); - - loopRes <-$ compFold _ (foldBody_option _ (ESPADA_TSetSetup_Sim_wLoop)) (Some (nil, initFree)) (zip qTags q_ts); - loopRes <-$ compFold _ (foldBody_option _ (fun acc e => randomTSetEntry acc)) loopRes (allNatsLt (minus (bigN t) (length (flatten q_ts)))); - tSet <- getTSet_opt loopRes; - inds <- map (fun x => lookupIndex (EqDec_dec _) (removeDups _ q) x O) q; - tags <- map (fun i => nth i qTags (Vector.const false lambda)) inds; - A2 s_A (tSet, tags). - - (* The previous game is identical to the ideal functionality with the ESPADA TSet simulator and leakage function. *) - - -End ESPADA_TSet_Games. \ No newline at end of file diff --git a/src/ESPADA/ESPADA_TSet_Secure.v b/src/ESPADA/ESPADA_TSet_Secure.v deleted file mode 100644 index 8a4c833..0000000 --- a/src/ESPADA/ESPADA_TSet_Secure.v +++ /dev/null @@ -1,1037 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import RepeatCore. -Require Import CompFold. -Require Import Permutation. -Require Import RndListElem. -Require Import Array. -Require Import PRF. - -Require Import TSet. -Require Import ESPADA_TSet_Once_Games. -Require Import ESPADA_TSet_Once. - -(* Use the "one-trial TSet" correctness proof along with the arguments in RepeatCore to get security of the full ESPADA TSet scheme. It would probably be better to prove this generically for any TSet scheme with a setup routine that may fail. *) - -Section ESPADA_TSet_secure. - - Variable lambda : posnat. - Variable A_State : Set. - Hypothesis A_State_EqDec : EqDec A_State. - Variable tSize bigB bigS : posnat. - - Variable F : Bvector lambda -> nat -> nat * Bvector lambda * Bvector (S lambda). - Variable F_bar : Bvector lambda -> Blist -> Bvector lambda. - - Hypothesis F_b_correct : - forall k b i v1 v2, - (b, v1, v2) = F k i -> - b < bigB. - - Variable A1 : Comp ((T lambda * list Blist) * A_State). - Variable A2 : A_State -> TSet lambda * list (Bvector lambda) -> Comp bool. - - Hypothesis A1_t_correct : - forall s_A t qs, - In ((t, qs), s_A) (getSupport A1) -> - bigN t = tSize. - - Hypothesis A1_wf : well_formed_comp A1. - Hypothesis A2_wf : forall s_A p, well_formed_comp (A2 s_A p). - - Definition Repeat_A2 (s_A : A_State)(p : (option (TSet lambda)) * list (Bvector lambda)) := - [tSet_opt, ts] <-2 p; - tSet <- - match tSet_opt with - | None => nil - | Some t => t - end; - A2 s_A (tSet, ts). - - (* numTrials is the number of attempts to obtain a valid TSet in an intermediate argument. It shows up as part of an adversary against the PRF that must be efficient, so it must be polynomial in the security parameter. It would be better if we could combine this with the success probability somehow. *) - Variable numTrials : posnat. - - Variable setupFailRate : Rat. - (* In order for this proof to be meaningful, the failure rate must be significantly less than one. But in order to complete the proof we require that it be not equal to 1. *) - Hypothesis setupFailRate_lt_1 : - ~ (1 <= setupFailRate). - - Definition setupSuccess(p : option (TSet lambda * FreeList) * Bvector lambda) := - if (fst p) then true else false. - Hypothesis setupFailRate_correct : - forall t qs s_A, - In (t, qs, s_A) (getSupport A1) -> - Pr[p <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; ret (negb (setupSuccess p))] <= setupFailRate. - - Definition simSetupSuccess(p : option (TSet lambda * FreeList) * list (Bvector lambda)) := - if (fst p) then true else false. - - Hypothesis simFailRate_correct : - forall t qs s_A, - In (t, qs, s_A) (getSupport A1) -> - Pr[p <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda t qs) (map (arrayLookupList _ t) qs); ret (negb (simSetupSuccess p))] <= setupFailRate. - - Theorem ESPADA_TSetSetup_Some_in_support : - forall t qs s_A, - (In (t, qs, s_A) (getSupport A1)) -> - exists k p, - In (Some p, k) (getSupport (ESPADA_TSetSetup_trial bigB bigS F F_bar t)). - - intuition. - - assert (~(Pr[p <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - ret (setupSuccess p) ] == 0)). - intuition. - - assert ( - Pr - [p <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; ret negb (setupSuccess p) ] + -Pr - [p <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; ret setupSuccess p ] == 1 -> False). - rewrite H0. - rewrite <- ratAdd_0_r. - intuition. - assert (1 <= setupFailRate). - rewrite <- H1. - eauto. - intuition. - - eapply H1. - - Lemma pred_both_sum_1 : - forall (A : Set)(c : Comp A)(P : A -> bool), - well_formed_comp c -> - Pr[a <-$ c; ret (P a)] + Pr[a <-$ c; ret negb (P a)] == 1. - - intuition. - simpl. - rewrite (sumList_filter_partition P). - rewrite ratAdd_comm. - rewrite (sumList_filter_partition P). - - assert (sumList (filter P (getSupport c)) - (fun b : A => - evalDist c b * - (if EqDec_dec bool_EqDec (negb (P b)) true then 1 else 0)) == 0). - eapply sumList_0. - intuition. - apply filter_In in H0. - intuition. - destruct (EqDec_dec bool_EqDec (negb (P a)) true). - destruct (P a); simpl in *; discriminate. - eapply ratMult_0_r. - rewrite H0. - clear H0. - rewrite <- ratAdd_0_l. - - assert (sumList (filter (fun a : A => negb (P a)) (getSupport c)) - (fun b : A => - evalDist c b * (if EqDec_dec bool_EqDec (P b) true then 1 else 0)) == 0). - eapply sumList_0. - intuition. - apply filter_In in H0. - intuition. - destruct (EqDec_dec bool_EqDec (P a) true ). - destruct (P a); simpl in *; discriminate. - eapply ratMult_0_r. - rewrite H0. - clear H0. - rewrite <- ratAdd_0_r. - - eapply eqRat_trans. - eapply ratAdd_eqRat_compat. - eapply sumList_body_eq. - intuition. - apply filter_In in H0. - intuition. - destruct (EqDec_dec bool_EqDec (negb (P a)) true). - eapply ratMult_1_r. - congruence. - - eapply sumList_body_eq; intuition. - apply filter_In in H0. - intuition. - destruct (EqDec_dec bool_EqDec (P a) true). - eapply ratMult_1_r. - congruence. - rewrite ratAdd_comm. - rewrite <- sumList_filter_partition. - eapply evalDist_lossless. - intuition. - Qed. - - rewrite ratAdd_comm. - eapply pred_both_sum_1. - - Theorem ESPADA_TSetSetup_trial_wf : - forall t, - well_formed_comp (ESPADA_TSetSetup_trial bigB bigS F F_bar t). - - intuition. - unfold ESPADA_TSetSetup_trial. - unfold ESPADA_TSet.ESPADA_TSetSetup_trial. - wftac. - eapply compFold_wf; intuition. - unfold ESPADA_TSetSetup_wLoop, ESPADA_TSet.ESPADA_TSetSetup_wLoop, foldBody_option. - destruct a; wftac. - destruct p. - eapply compFold_wf; intuition. - destruct a; wftac. - unfold ESPADA_TSet.ESPADA_TSetSetup_tLoop. - destruct p. - destruct (F (F_bar b b0) a0). - destruct p. - comp_simp. - unfold maybeBindComp. - wftac. - eapply rndListElem_wf. - intuition. - destruct b4; wftac. - - Qed. - - eapply ESPADA_TSetSetup_trial_wf. - - Lemma bind_ret_nz_exists : - forall (A : Set)(c : Comp A)(f : A -> bool), - (~(Pr[a <-$ c; ret (f a)] == 0)) -> - exists a, In a (getSupport c) /\ f a = true. - - intuition. - simpl in H. - apply sumList_nz in H. - destruct H. - intuition. - econstructor. - intuition. - eauto. - destruct (EqDec_dec bool_EqDec (f x) true); intuition. - exfalso. - eapply H1. - eapply ratMult_0_r. - Qed. - - apply bind_ret_nz_exists in H0. - destruct H0. - intuition. - destruct x. - unfold setupSuccess in *. - destruct o. - simpl in H2. - econstructor. - econstructor. - eauto. - - simpl in H2. - discriminate. - Qed. - - Theorem ESPADA_SimSetup_Some_in_support : - forall t qs s_A, - (In (t, qs, s_A) (getSupport A1)) -> - exists r p, - In (Some p, r) (getSupport (ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda t qs) (map (arrayLookupList _ t) qs))). - - intuition. - - assert (~(Pr[p <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda t qs) (map (arrayLookupList _ t) qs); - ret (simSetupSuccess p) ] == 0)). - intuition. - - assert (Pr - [p <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda t qs) (map (arrayLookupList _ t) qs); - ret negb (simSetupSuccess p) ] + - Pr - [p <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda t qs) (map (arrayLookupList _ t) qs); - ret simSetupSuccess p ] == 1 -> False). - - rewrite H0. - rewrite <- ratAdd_0_r. - intuition. - assert (1 <= setupFailRate). - rewrite <- H1. - eauto. - intuition. - - eapply H1. - - rewrite ratAdd_comm. - eapply pred_both_sum_1. - - Lemma ESPADA_TSet_Sim_trial_wf : - forall a0 l0 qs, - well_formed_comp (ESPADA_TSet_Sim_trial lambda bigB bigS F (@L_T lambda a0 qs) l0). - - intuition. - unfold ESPADA_TSet_Sim_trial; wftac. - eapply compMap_wf; wftac. - - eapply compFold_wf; intuition. - unfold ESPADA_TSetSetup_Sim_wLoop, foldBody_option. - destruct a; wftac. - destruct p. - eapply compFold_wf; intuition. - destruct a; wftac. - unfold ESPADA_TSetSetup_tLoop. - unfold ESPADA_TSet.ESPADA_TSetSetup_tLoop. - destruct p. - destruct (F a1 a2). - destruct p. - comp_simp. - unfold maybeBindComp. - wftac. - eapply rndListElem_wf. - intuition. - destruct b4; wftac. - eapply compFold_wf; intuition. - unfold foldBody_option. - destruct a; wftac. - unfold randomTSetEntry. - wftac. - intuition. - unfold maybeBindComp. - wftac. - eapply rndListElem_wf. - intuition. - destruct b6; wftac. - - Qed. - - eapply ESPADA_TSet_Sim_trial_wf. - - apply bind_ret_nz_exists in H0. - destruct H0. - intuition. - destruct x. - unfold simSetupSuccess in *. - destruct o. - simpl in H2. - econstructor. - econstructor. - eauto. - - simpl in H2. - discriminate. - - Qed. - - Theorem TSet_RepeatCore_Real_eq : - Pr [TSetReal (Bvector_EqDec lambda) (ESPADA_TSetSetup bigB bigS F F_bar) - (ESPADA_TSetGetTag lambda F_bar) A1 A2 ] == - Pr[RepeatCore_G - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) A1 Repeat_A2 - (Repeat_Real bigB bigS F F_bar) ]. - - unfold TSetReal, RepeatCore_G, Repeat_A2, Repeat_Real, ESPADA_TSetSetup_once, ESPADA_TSetSetup. - - comp_skip; comp_simp. - unfold TSet.A, fst. - intuition. - - assert (Pr - [b1 <-$ - Repeat - (z <-$ - (z <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - [res, k_T]<-2 z; ret (getTSet_opt res, k_T)); - [tSet, k_T]<-2 z; - tags <-$ - compMap (Bvector_EqDec lambda) (ESPADA_TSetGetTag lambda F_bar k_T) - l ; ret (tSet, tags)) - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false); - - - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ]== - Pr - [b1 <-$ - Repeat - (z <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - tags <-$ - compMap (Bvector_EqDec lambda) (ESPADA_TSetGetTag lambda F_bar (snd z)) - l; ret (getTSet_opt (fst z), tags)) - (fun p => - if fst p then true else false); - - - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ]). - - comp_skip. - - case_eq a0; intuition. - destruct (in_dec (EqDec_dec _ ) (Some t0, b) (getSupport (z <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - tags <-$ - compMap (Bvector_EqDec lambda) - (ESPADA_TSetGetTag lambda F_bar (snd z)) - l; - ret (getTSet_opt (fst z), tags)))). - - eapply evalDist_Repeat_eq; intuition. - comp_inline_l. - comp_skip; comp_simp. - comp_skip. - simpl. - eapply eqRat_refl. - simpl. - intuition. - - subst. - repeat simp_in_support. - destruct x. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - remember ((getTSet_opt o, b)) as v. - destruct v. - pairInv. - eapply getSupport_In_Seq. - eapply H2. - simpl. - left. - f_equal. - intuition. - - repeat simp_in_support. - eapply eqRat_trans. - eapply (@sumList_permutation _ _ _ (filter - (fun - p : option (ESPADA_TSet_Once_Games.TSet lambda) * list (Bvector lambda) => - if fst p then true else false) - (getSupport - (z <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - tags <-$ - compMap (Bvector_EqDec lambda) - (ESPADA_TSetGetTag lambda F_bar (snd z)) - l; - ret (getTSet_opt (fst z), tags))))). - - eapply NoDup_Permutation; intuition. - eapply filter_NoDup. - eapply getSupport_NoDup. - eapply filter_NoDup. - eapply getSupport_NoDup. - apply filter_In in H0; intuition. - repeat simp_in_support. - destruct x1. - repeat simp_in_support. - destruct x2. - repeat simp_in_support. - destruct x. - simpl in H4. - simpl in H5. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eauto. - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - apply filter_In in H0; intuition. - repeat simp_in_support. - destruct x1. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply getSupport_In_Seq. - eauto. - comp_simp. - simpl. - intuition. - comp_simp. - - eapply getSupport_In_Seq. - eauto. - simpl. - intuition. - - eapply sumList_body_eq; intuition. - comp_inline_l. - comp_skip; comp_simp. - comp_skip. - simpl. - intuition. - simpl. - intuition. - - repeat rewrite evalDistRepeat_sup_0; intuition; - eapply getSupport_not_In_evalDist; intuition; - repeat simp_in_support. - destruct x. - repeat simp_in_support. - destruct x0. - repeat simp_in_support. - eapply n. - eapply getSupport_In_Seq. - eauto. - eapply getSupport_In_Seq. - eauto. - simpl. - rewrite H1. - intuition. - - repeat rewrite evalDistRepeat_pred_0. - intuition. - simpl. intuition. - simpl. intuition. - - rewrite H0. - clear H0. - - assert ( - Pr - [b1 <-$ - Repeat - (z <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; - tags <-$ - compMap (Bvector_EqDec lambda) - (ESPADA_TSetGetTag lambda F_bar (snd z)) l; - ret (getTSet_opt (fst z), tags)) - (fun p : option (ESPADA_TSet_Once_Games.TSet lambda) * list (Bvector lambda) => - if fst p then true else false); - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ] == - Pr - [b1 <-$ - (z <-$ Repeat - (ESPADA_TSetSetup_trial bigB bigS F F_bar t) - (fun p => if fst p then true else false); - tags <-$ compMap (Bvector_EqDec lambda) - (ESPADA_TSetGetTag lambda F_bar (snd z)) l; - ret (getTSet_opt (fst z), tags)); - - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ]). - - comp_skip. - symmetry. - eapply (@repeat_fission' _ _ _ _) ; intuition. - - apply ESPADA_TSetSetup_trial_wf. - edestruct ESPADA_TSetSetup_Some_in_support . - eauto. - destruct H0. - econstructor. - eapply filter_In; intuition. - eapply H0. - simpl. - trivial. - - wftac. - eapply compMap_wf; intuition. - unfold ESPADA_TSetGetTag. - wftac. - - repeat simp_in_support. - unfold getTSet_opt. - simpl. - destruct a2; trivial. - - rewrite H0. - clear H0. - inline_first. - comp_skip. - eapply eqRat_refl. - - comp_simp. - inline_first. - comp_skip. - simpl. - eapply eqRat_refl. - comp_simp. - simpl. - destruct o; simpl; intuition. - reflexivity. - reflexivity. - Qed. - - Theorem TSet_RepeatCore_Ideal_eq : - Pr [TSetIdeal (L_T (lambda:=lambda)) A1 A2 - (ESPADA_TSet_Sim lambda bigB bigS F) ] == - Pr - [RepeatCore_G - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) A1 Repeat_A2 - (Repeat_Ideal lambda bigB bigS F) ]. - - unfold TSetIdeal, RepeatCore_G, Repeat_Ideal, ESPADA_TSet_Sim, ESPADA_TSet_Sim_once, Repeat_A2, L_T. - - comp_skip. - destruct x. - destruct p. - - assert ( Pr - [b1 <-$ - Repeat - (z <-$ - ESPADA_TSet_Sim_trial lambda bigB bigS F - (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l); - [trialRes, tags]<-2 z; ret (getTSet_opt trialRes, tags)) - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false); - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ] == - Pr - [b1 <-$ - (p <-$ Repeat - (ESPADA_TSet_Sim_trial lambda bigB bigS F - (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l)) - (fun p => if fst p then true else false); - ret (getTSet_opt (fst p), (snd p))); - [tSet_opt, ts]<-2 b1; - A2 a (match tSet_opt with - | Some t => t - | None => nil - end, ts) ] -). - - comp_skip. - symmetry. - - assert (evalDist - (p <-$ - Repeat (ESPADA_TSet_Sim_trial lambda bigB bigS F - (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l)) - (fun - p : option (ESPADA_TSet_Once_Games.TSet lambda * FreeList) * - list (Bvector lambda) => if fst p then true else false); - ret (getTSet_opt (fst p), snd p)) (a0, b) == - evalDist - (Repeat (z <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F - (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l); - ret (getTSet_opt (fst z), snd z)) - (fun p => if fst p then true else false)) (a0, b)). - eapply (@repeat_fission' _ _ _ _) ; intuition. - - eapply ESPADA_TSet_Sim_trial_wf. - edestruct (ESPADA_SimSetup_Some_in_support). - eauto. - destruct H0. - econstructor. - eapply filter_In; intuition. - eapply H0. - trivial. - wftac. - - repeat simp_in_support. - simpl. - unfold getTSet_opt. - destruct a2; trivial. - - rewrite H0. - clear H0. - - case_eq a0; intuition. - destruct (in_dec (EqDec_dec _) (Some t0, b) (getSupport (z <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l); - ret (getTSet_opt (fst z), snd z)))). - - eapply evalDist_Repeat_eq. - comp_skip; comp_simp. - simpl. - intuition. - trivial. - - repeat simp_in_support. - eapply filter_In; intuition. - eapply getSupport_In_Seq; eauto. - simpl. - left. - f_equal. - intuition. - - eapply eqRat_trans. - eapply (@sumList_permutation _ _ _ - (filter - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) - (getSupport - (z <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F - (bigN t, - map - (fun x : Blist => - CompFold.lookupIndex (EqDec_dec (list_EqDec bool_EqDec)) - (removeDups (list_EqDec bool_EqDec) l) x 0) l) - (map (arrayLookupList (list_EqDec bool_EqDec) t) l); - [trialRes, tags]<-2 z; ret (getTSet_opt trialRes, tags))))). - eapply NoDup_Permutation; intuition. - eapply filter_NoDup. - eapply getSupport_NoDup. - eapply filter_NoDup. - eapply getSupport_NoDup. - apply filter_In in H1. - intuition. - repeat simp_in_support. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply H2. - destruct x0. - destruct x. - simpl. - intuition. - apply filter_In in H1. - repeat simp_in_support. - destruct x0. - simp_in_support. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply H2. - destruct x. - simpl. - intuition. - - eapply sumList_body_eq; intuition. - comp_skip. - comp_simp. - simpl. - intuition. - - repeat rewrite evalDistRepeat_sup_0; intuition; - eapply getSupport_not_In_evalDist; intuition; - repeat simp_in_support. - destruct x. - eapply n. - eapply getSupport_In_Seq; eauto. - - repeat rewrite evalDistRepeat_pred_0; intuition. - rewrite H0. - clear H0. - inline_first. - comp_skip. - - eapply eqRat_refl. - - comp_simp. - simpl. - destruct o; simpl; eapply eqRat_refl. - - Qed. - - Theorem ESPADA_TSet_eq_RepeatCore : - (| Pr[TSetReal _ (@ESPADA_TSetSetup lambda bigB bigS F F_bar) (@ESPADA_TSetGetTag lambda F_bar) A1 A2] - - Pr[TSetIdeal (@L_T lambda) A1 A2 (@ESPADA_TSet_Sim lambda bigB bigS F)] |) == - RepeatCore_Adv (fun (p : option (TSet lambda) * list (Bvector lambda)) => if (fst p) then true else false) (Repeat_Real bigB bigS F F_bar) (Repeat_Ideal lambda bigB bigS F) A1 Repeat_A2. - - intuition. - - unfold RepeatCore_Adv. - - rewrite TSet_RepeatCore_Real_eq. - rewrite TSet_RepeatCore_Ideal_eq. - intuition. - Qed. - - - Theorem Repeat_Real_fail_eq : - forall t q, - Pr [b0 <-$ Repeat_Real bigB bigS F F_bar (t, q); ret negb (if fst b0 then true else false) ] - == - Pr[p <-$ ESPADA_TSetSetup_trial bigB bigS F F_bar t; ret (negb (setupSuccess p))]. - - intuition. - unfold Repeat_Real, ESPADA_TSetSetup_once. - comp_inline_l. - comp_inline_l. - comp_skip. - comp_simp. - unfold setupSuccess. - inline_first. - comp_irr_l. - eapply compMap_wf; intuition. - unfold ESPADA_TSetGetTag. - wftac. - - destruct o; - simpl; intuition. - Qed. - - Theorem Repeat_Ideal_fail_eq : - forall t q, - Pr[b0 <-$ Repeat_Ideal lambda bigB bigS F (t, q); ret negb (if fst b0 then true else false) ] - == - Pr[p <-$ ESPADA_TSet_Sim_trial lambda bigB bigS F (L_T t q) (map (arrayLookupList _ t) q); ret (negb (simSetupSuccess p))]. - - intuition. - unfold Repeat_Ideal. - unfold ESPADA_TSet_Sim_once, L_T. - comp_inline_l. - comp_skip. - comp_simp. - - unfold simSetupSuccess. - unfold fst. - simpl. - destruct o; - simpl; - intuition. - - Qed. - - Definition Repeat_PRF_A1 := - (ESPADA_TSet_PRF_A1 _ - (B1 _ _ A1)). - - Definition ESPADA_TSet_PRF_A2 := - ESPADA_TSet_PRF_A2 bigB bigS F. - - Definition Repeat_PRF_A2 := - (ESPADA_TSet_PRF_A2 - (B2 _ - (Repeat_Real bigB bigS F F_bar) (Repeat_Ideal lambda bigB bigS F) - (DM_RC_B2 - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) Repeat_A2) numTrials) - ). - - Definition ESPADA_TSet_IPRF_A1 := - ESPADA_TSet_IPRF_A1 bigB bigS F. - - Definition Repeat_IPRF_A1 := - (ESPADA_TSet_IPRF_A1 _ - (B1 _ _ A1)). - - - Definition Repeat_IPRF_A2 := - (ESPADA_TSet_IPRF_A2 - (B2 _ - (Repeat_Real bigB bigS F F_bar) (Repeat_Ideal lambda bigB bigS F) - (DM_RC_B2 - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) Repeat_A2) numTrials)). - - Lemma Repeat_Ideal_wf : - forall a : T lambda * list Blist, - well_formed_comp (Repeat_Ideal lambda bigB bigS F a). - - intuition. - unfold Repeat_Ideal. - unfold ESPADA_TSet_Sim_once. - eapply well_formed_Bind. - eapply ESPADA_TSet_Sim_trial_wf. - intuition. - wftac. - - Qed. - - Lemma Repeat_Real_wf : - forall a : T lambda * list Blist, - well_formed_comp (Repeat_Real bigB bigS F F_bar a). - - intuition. - unfold Repeat_Real. - unfold ESPADA_TSetSetup_once. - wftac. - eapply ESPADA_TSetSetup_trial_wf. - eapply compMap_wf; intuition. - unfold ESPADA_TSetGetTag. - wftac. - Qed. - - Lemma Repeat_Ideal_Repeat_wf : - forall t qs s_A, - In (t, qs, s_A) (getSupport A1) -> - exists b : option (TSet lambda) * list (Bvector lambda), - In b - (filter - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) - (getSupport (Repeat_Ideal lambda bigB bigS F (t, qs)))). - - intuition. - edestruct ESPADA_SimSetup_Some_in_support. - eauto. - destruct H0. - - econstructor. - unfold Repeat_Ideal. - unfold ESPADA_TSet_Sim_once, L_T. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply H0. - comp_simp. - simpl. - intuition. - trivial. - - Qed. - - Lemma Repeat_Real_Repeat_wf : - forall t qs s_A, - In (t, qs, s_A) (getSupport A1) -> - exists b : option (TSet lambda) * list (Bvector lambda), - In b - (filter - (fun p : option (TSet lambda) * list (Bvector lambda) => - if fst p then true else false) - (getSupport (Repeat_Real bigB bigS F F_bar (t, qs)))). - - intuition. - edestruct ESPADA_TSetSetup_Some_in_support. - eauto. - destruct H0. - - edestruct (@well_formed_val_exists _ (compMap (Bvector_EqDec lambda) (ESPADA_TSetGetTag lambda F_bar x) - qs)). - eapply compMap_wf; intuition. - unfold ESPADA_TSetGetTag. - wftac. - - econstructor. - unfold Repeat_Real. - unfold ESPADA_TSetSetup_once. - eapply filter_In; intuition. - eapply getSupport_In_Seq. - eapply getSupport_In_Seq. - eauto. - - comp_simp. - simpl. - intuition. - comp_simp. - - eapply getSupport_In_Seq. - eapply H1. - simpl. - intuition. - simpl. - trivial. - Qed. - - Variable maxKeywords : nat. - Hypothesis maxKeywords_correct : - (forall db q s_A, - In (db, q, s_A) (getSupport A1) -> - length (combineKeywords (removeDups _ q) (toW db)) <= maxKeywords)%nat. - - Variable PRF_NA_Adv : Rat. - Hypothesis PRF_NA_Adv_correct_1 : - (PRF_NA_Advantage (Rnd lambda) ({ 0 , 1 }^lambda) F_bar _ _ Repeat_PRF_A1 Repeat_PRF_A2) <= PRF_NA_Adv. - - Hypothesis PRF_NA_Adv_correct_2: - forall i, - PRF_NA_Advantage (Rnd lambda) (@RndF_range lambda bigB) F _ _ (Hybrid.B1 nil _ _ Repeat_IPRF_A1 i) - (Hybrid.B2 - (fun lsD => k <-$ {0, 1}^lambda; ret (map (F k) lsD)) - (fun lsD => [lsR, _] <-$2 oracleMap _ _ (RndR_func (@RndF_range lambda bigB) _) nil lsD; ret lsR) - _ Repeat_IPRF_A2) - <= PRF_NA_Adv. - - Theorem ESPADA_TSet_secure : - | Pr[TSetReal _ (@ESPADA_TSetSetup lambda bigB bigS F F_bar) (@ESPADA_TSetGetTag lambda F_bar) A1 A2] - - Pr[TSetIdeal (@L_T lambda) A1 A2 (@ESPADA_TSet_Sim lambda bigB bigS F)] | - <= numTrials / 1 * - (PRF_NA_Adv + - (maxKeywords / 1) * - PRF_NA_Adv) + - expRat setupFailRate numTrials + expRat setupFailRate numTrials. - - rewrite ESPADA_TSet_eq_RepeatCore. - rewrite DistMult_impl_RepeatCore. - rewrite DistSingle_impl_Mult. - rewrite ESPADA_TSet_DistSingle_secure. - - eapply ratAdd_leRat_compat; [idtac | eapply leRat_refl]. - eapply ratAdd_leRat_compat; [idtac | eapply leRat_refl]. - eapply ratMult_leRat_compat; intuition. - unfold Repeat_PRF_A1, Repeat_PRF_A2, Repeat_IPRF_A1, Repeat_IPRF_A2. - eapply leRat_refl. - - apply (fun n => (pos 1, pos 1)). - - intuition. - - intuition. - unfold B1 in *. - repeat simp_in_support. - destruct x. - repeat simp_in_support. - eauto. - - eauto. - eauto. - - intuition. - apply Repeat_Ideal_wf. - intuition. - intuition. - - intuition. - unfold Repeat_A2. - destruct a; wftac. - - intuition. - apply Repeat_Real_wf. - - intuition. - eapply Repeat_Real_Repeat_wf. - eauto. - - intuition. - apply Repeat_Ideal_wf. - - intuition. - eapply Repeat_Ideal_Repeat_wf. - eauto. - - intuition. - rewrite Repeat_Real_fail_eq. - eauto. - - intuition. - rewrite Repeat_Ideal_fail_eq. - eauto. - Qed. - - Print Assumptions ESPADA_TSet_secure. - -End ESPADA_TSet_secure. \ No newline at end of file diff --git a/src/ESPADA/Makefile b/src/ESPADA/Makefile deleted file mode 100644 index 0032f2d..0000000 --- a/src/ESPADA/Makefile +++ /dev/null @@ -1,16 +0,0 @@ - -MODULES := TSet ESPADA_TSet ESPADA_TSet_Once_Games ESPADA_TSet_Once ESPADA_TSet_Secure SSE ESPADA_TSet_Correct_Once_Games ESPADA_TSet_Correct_Once ESPADA_TSet_Correct ESPADA_SSE_SKS ESPADA_SSE_SKS_Secure ESPADA_SSE_SKS_Secure_tacs - -VS := $(MODULES:%=%.v) - -.PHONY: coq clean - -coq: Makefile.coq - $(MAKE) -f Makefile.coq - -Makefile.coq: Makefile $(VS) - coq_makefile -I "../FCF" $(VS) -o Makefile.coq - -clean:: Makefile.coq - $(MAKE) -f Makefile.coq clean - rm -f Makefile.coq diff --git a/src/ESPADA/SSE.v b/src/ESPADA/SSE.v deleted file mode 100644 index 769d31d..0000000 --- a/src/ESPADA/SSE.v +++ /dev/null @@ -1,56 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) -Set Implicit Arguments. - -Require Import FCF. -Require Import CompFold. -Require Import Array. - -Local Open Scope type_scope. -Local Open Scope list_scope. - -Section SSE. - - Variable lambda : nat. - Definition Identifier := Bvector lambda. - Definition Keyword := Blist. - Definition DB_Entry := (Identifier * list Keyword). - Definition DB := list DB_Entry. - Variable Query : Set. - - Variable SecretKey : Set. - Variable EDB : Set. - Variable EDBSetup : DB -> Comp (SecretKey * EDB). - Variable SearchTranscript : Set. - Hypothesis SearchTranscript_EqDec : EqDec SearchTranscript. - Variable SearchProtocol : EDB -> SecretKey -> Query -> Comp (list Identifier * SearchTranscript). - - Section SSE_Security_NA. - - Variable A_State : Set. - Variable A1 : Comp (DB * list Query * A_State). - Variable A2 : A_State -> EDB -> list SearchTranscript -> Comp bool. - - Variable Leakage : Set. - Variable L : DB -> list Query -> Comp Leakage. - Variable Sim : Leakage -> Comp (EDB * list SearchTranscript). - - Definition SSE_Sec_NA_Real := - [db, q, s_A] <-$3 A1; - [k, edb] <-$2 EDBSetup db; - ls <-$ compMap _ (SearchProtocol edb k) q; - A2 s_A edb (snd (split ls)). - - Definition SSE_Sec_NA_Ideal := - [db, q, s_A] <-$3 A1; - leak <-$ L db q; - [edb, t] <-$2 Sim leak; - A2 s_A edb t. - - Definition SSE_NA_Advantage := - | Pr[SSE_Sec_NA_Real] - Pr[SSE_Sec_NA_Ideal] |. - - End SSE_Security_NA. - -End SSE. \ No newline at end of file diff --git a/src/ESPADA/TSet.v b/src/ESPADA/TSet.v deleted file mode 100644 index 8b91f09..0000000 --- a/src/ESPADA/TSet.v +++ /dev/null @@ -1,79 +0,0 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * - * in the LICENSE file at the root of the source tree. *) - -Set Implicit Arguments. - -Require Import FCF. -Require Import CompFold. -Require Import Array. - -Local Open Scope array_scope. - -Section TSet. - - Variable lambda : posnat. (* security parameter *) - Definition T := list (Blist * (list (Bvector lambda))). - Variable TSet : Set. - Variable Tag : Set. - Variable SecretKey : Set. - Hypothesis Tag_EqDec : EqDec Tag. - - Variable TSetSetup : T -> Comp (TSet * SecretKey). - Variable TSetGetTag : SecretKey -> Blist -> Comp Tag. - Variable TSetRetrieve : TSet -> Tag -> (list (Bvector lambda)). - - Section TSetCorrectness. - - Variable A : Comp (T * list Blist). - - Definition AdvCor_G := - [t, q] <-$2 A; - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ compMap _ (TSetGetTag k_T) q; - t_w <- map (TSetRetrieve tSet) tags; - t_w_correct <- map (arrayLookupList _ t) q; - ret (negb (eqb t_w t_w_correct)). - - Definition AdvCor := Pr[AdvCor_G]. - - (* AdvCor should be small for a correct TSet. *) - - End TSetCorrectness. - - (* Security against a non-adaptive adversary. In this definition, the adversary must choose the list of queries before seeing the encrypted structure. *) - Section TSetSecurity. - - Variable Leakage : Set. - Variable L : T -> list Blist -> Leakage. - - Variable A_State : Set. - Variable A1 : Comp (T * list Blist * A_State). - Variable A2 : A_State -> TSet * list Tag -> Comp bool. - Definition A := (A1, A2). - - (* For simplicity, we remove duplicates from the list of queries before processing it. *) - - Definition TSetReal := - [t, q, s_A] <-$3 A1; - [tSet, k_T] <-$2 TSetSetup t; - tags <-$ foreach (x in q) (TSetGetTag k_T x); - A2 s_A (tSet, tags). - - - Variable Sim : Leakage -> list (list (Bvector lambda)) -> Comp (TSet * list Tag). - - Definition lookupAnswers (t : Array Blist (list (Bvector lambda))) (q : Blist) := - arrayLookupList _ t q. - - Definition TSetIdeal := - [t, q, s_A] <-$3 A1; - T_qs <- foreach (x in q) (lookupAnswers t x); - [tSet, tags] <-$2 Sim (L t q) T_qs; - A2 s_A (tSet, tags). - - Definition TSetAdvantage := | Pr[TSetReal] - Pr[TSetIdeal] |. - - End TSetSecurity. - -End TSet. \ No newline at end of file diff --git a/src/FCF/Blist.v b/src/FCF/Blist.v index b18f048..e177098 100644 --- a/src/FCF/Blist.v +++ b/src/FCF/Blist.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* Lists of booleans and related theory. *) @@ -216,13 +215,13 @@ Theorem shiftOut_S_None : forall (n : nat)(s s1 : Blist)(v1 : Bvector 1), shiftOut s 1 = Some (v1, s1) -> shiftOut s1 n = None -> shiftOut s (S n) = None. -Admitted. +Abort. Theorem shiftOut_1_None : forall (n1 n2 : nat)(s : Blist), shiftOut s n1 = None -> n2 >= n1 -> shiftOut s n2 = None. -Admitted. +Abort. (* Todo : we need a general theorem that covers these sorts of facts. Something like: shiftOut s n1 = Some(_, s1) -> @@ -1116,4 +1115,4 @@ Lemma bvToNat_natToBv_eq : forall n (v : Bvector n) k, rewrite <- H. symmetry. apply natToBv_bvToNat_inverse. -Qed. \ No newline at end of file +Qed. diff --git a/src/FCF/Encryption_2W.v b/src/FCF/Broken/Encryption_2W.v similarity index 96% rename from src/FCF/Encryption_2W.v rename to src/FCF/Broken/Encryption_2W.v index 2e50ac4..6bde6f5 100644 --- a/src/FCF/Encryption_2W.v +++ b/src/FCF/Broken/Encryption_2W.v @@ -1,9 +1,10 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* A "two worlds" style of definition for encryption and a proof that it is equivalent to the standard definition. *) +(* + Set Implicit Arguments. Require Import Crypto. @@ -148,4 +149,5 @@ Section Encryption_SecretKey_2W. Qed. -End Encryption_SecretKey_2W. \ No newline at end of file +End Encryption_SecretKey_2W. +*) diff --git a/src/FCF/ListHybrid.v b/src/FCF/Broken/ListHybrid.v similarity index 59% rename from src/FCF/ListHybrid.v rename to src/FCF/Broken/ListHybrid.v index 04e1649..335b55e 100644 --- a/src/FCF/ListHybrid.v +++ b/src/FCF/Broken/ListHybrid.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) Require Import FCF. @@ -132,150 +131,173 @@ Section ListHybrid. rewrite map_app. simpl. - - + (* eapply eqRat_impl_leRat. rewrite ratS_num. rewrite ratMult_distrib_r. eapply ratAdd_eqRat_compat; intuition. symmetry. eapply ratMult_1_l. + *) + +(* 1 subgoal, subgoal 1 (ID 390) *) + +(* A, B, C, S_A : Set *) +(* eqdb : EqDec B *) +(* defA : A *) +(* A1 : Comp (list A * S_A) *) +(* A2, A2' : S_A -> A -> Comp B *) +(* A3 : S_A -> list B -> Comp C *) +(* A1_wf : well_formed_comp A1 *) +(* A2_wf : forall (x : S_A) (y : A), well_formed_comp (A2 x y) *) +(* A2'_wf : forall (x : S_A) (y : A), well_formed_comp (A2' x y) *) +(* numA : nat *) +(* numA_correct : forall (lsa : list A) (s : S_A), *) +(* In (lsa, s) (getSupport A1) -> length lsa = numA *) +(* F : nat -> Rat *) +(* H : forall (n : nat) (c : C), *) +(* | evalDist (LH_Gn n) c - evalDist (LH_Gn (S n)) c | <= F n *) +(* x : nat *) +(* IHx : forall (n : nat) (c : C), *) +(* | evalDist (LH_Gn n) c - evalDist (LH_Gn (n + x)) c | <= *) +(* sumList (map (Init.Nat.add n) (allNatsLt x)) F *) +(* n : nat *) +(* c : C *) +(* ============================ *) +(* F n + sumList (map (fun m : nat => S (n + m)) (allNatsLt x)) F <= *) +(* sumList (map (Init.Nat.add n) (allNatsLt x) ++ (n + x)%nat :: nil) F *) + +(* (dependent evars: (printing disabled) ) *) + + Admitted. + + Theorem skipn_length_nil : + forall (A : Type)(ls : list A)(n : nat), + n >= length ls -> + skipn n ls = nil. + + induction ls; intuition; simpl in *. + destruct n; simpl in *; trivial. + + destruct n; simpl in *. + omega. + eapply IHls. + omega. Qed. - Theorem list_hybrid_close : - forall k c, - (forall n c, n < numA -> | (evalDist (LH_G_0 n) c) - (evalDist (LH_G_1 n) c) | <= k) -> - | (evalDist (LH_G0) c) - (evalDist (LH_G1) c) | <= (numA)/1 * k. - - intuition. + Theorem LH_Gn_0_equiv : + forall (n : nat) c, + n < numA -> + evalDist (LH_Gn n) c == evalDist (LH_G_0 n) c. - assert (evalDist (LH_G0) c == evalDist (LH_Gn 0) c). - unfold LH_G0, LH_Gn. + intuition. + unfold LH_Gn, LH_G_0. + comp_skip; comp_simp. comp_skip. - comp_simp. - simpl. - comp_simp. + + erewrite skipn_S. simpl. + inline_first. + comp_skip. + reflexivity. + inline_first. comp_skip. + erewrite numA_correct; intuition; eauto. + + Qed. - rewrite H0. - clear H0. + Theorem LH_Gn_1_equiv : + forall (n : nat) c, + n < numA -> + evalDist (LH_Gn (S n)) c == evalDist (LH_G_1 n) c. - assert (evalDist (LH_G1) c == evalDist (LH_Gn numA) c). - unfold LH_G1, LH_Gn. + intuition. + unfold LH_Gn, LH_G_1. comp_skip. comp_simp. - specialize (firstn_skipn (length l) l); intuition. - rewrite <- H1 at 1. - eapply eqRat_trans. - eapply evalDist_seq_eq. - intros. + rewrite (firstn_S _ defA). + + assert (evalDist (lsb1 <-$ (foreach (x in firstn n l ++ nth n l defA :: nil) A2' s x); + lsb2 <-$ (foreach (x in skipn (S n) l)A2 s x); A3 s (lsb1 ++ lsb2) ) c == + evalDist (lsb1 <-$ (r1 <-$ compMap _ (A2' s) (firstn n l); + r2 <-$ compMap _ (A2' s) (nth n l defA :: nil); + ret (r1 ++ r2)); + lsb2 <-$ (foreach (x in skipn (S n) l)A2 s x); A3 s (lsb1 ++ lsb2) ) c + ). + + comp_skip. eapply compMap_app. - intuition. - reflexivity. + match goal with [H1:_|- _ ] => rewrite H1; clear H1 end. + inline_first. + comp_skip. inline_first. - clear H1. comp_skip. - erewrite numA_correct; eauto. - intuition. inline_first. - Theorem skipn_length_nil : - forall (A : Type)(ls : list A)(n : nat), - n >= length ls -> - skipn n ls = nil. - - induction ls; intuition; simpl in *. - destruct n; simpl in *; trivial. - - destruct n; simpl in *. - omega. - eapply IHls. - omega. - Qed. - - repeat rewrite skipn_length_nil. + comp_skip. + reflexivity. + + rewrite <- app_assoc. + rewrite <- app_comm_cons. simpl. - comp_simp. intuition. erewrite numA_correct; intuition; eauto. - intuition. - - rewrite H0. - clear H0. - assert (numA = (0 + numA))%nat. - omega. - rewrite H0 at 1. - clear H0. - rewrite list_hybrid_close_h. - eapply leRat_refl. - intuition. + Qed. - Theorem LH_Gn_0_equiv : - forall (n : nat) c, - n < numA -> - evalDist (LH_Gn n) c == evalDist (LH_G_0 n) c. + Theorem list_hybrid_close : + forall k c, + (forall n c, n < numA -> | (evalDist (LH_G_0 n) c) - (evalDist (LH_G_1 n) c) | <= k) -> + | (evalDist (LH_G0) c) - (evalDist (LH_G1) c) | <= (numA)/1 * k. + pose proof @skipn_length_nil as skipn_length_nil. + pose proof @list_hybrid_close_h as list_hybrid_close_h. + + intros. - intuition. - unfold LH_Gn, LH_G_0. - comp_skip; comp_simp. + assert (evalDist (LH_G0) c == evalDist (LH_Gn 0) c) as HA. { + unfold LH_G0, LH_Gn. comp_skip. - - erewrite skipn_S. + comp_simp. + simpl. + comp_simp. simpl. - inline_first. - comp_skip. - reflexivity. - inline_first. comp_skip. - erewrite numA_correct; intuition; eauto. - - Qed. + } rewrite HA; clear HA. - Theorem LH_Gn_1_equiv : - forall (n : nat) c, - n < numA -> - evalDist (LH_Gn (S n)) c == evalDist (LH_G_1 n) c. - - intuition. - unfold LH_Gn, LH_G_1. + assert (evalDist (LH_G1) c == evalDist (LH_Gn numA) c) as HA. { + unfold LH_G1, LH_Gn. comp_skip. comp_simp. - rewrite (firstn_S _ defA). - - assert (evalDist (lsb1 <-$ (foreach (x in firstn n l ++ nth n l defA :: nil) A2' s x); - lsb2 <-$ (foreach (x in skipn (S n) l)A2 s x); A3 s (lsb1 ++ lsb2) ) c == - evalDist (lsb1 <-$ (r1 <-$ compMap _ (A2' s) (firstn n l); - r2 <-$ compMap _ (A2' s) (nth n l defA :: nil); - ret (r1 ++ r2)); - lsb2 <-$ (foreach (x in skipn (S n) l)A2 s x); A3 s (lsb1 ++ lsb2) ) c - ). - - comp_skip. + rewrite <- (firstn_skipn (length l) l) at 1 by intuition. + eapply eqRat_trans. + eapply evalDist_seq_eq. + intros. eapply compMap_app. - rewrite H1. - clear H1. - inline_first. - comp_skip. - inline_first. - comp_skip. + intuition. + reflexivity. inline_first. - comp_skip. + erewrite numA_correct; eauto. reflexivity. + inline_first. - rewrite <- app_assoc. - rewrite <- app_comm_cons. + replace numA with (length l) by (eapply numA_correct; eauto). + rewrite skipn_length_nil. simpl. + comp_simp. + reflexivity. intuition. - erewrite numA_correct; intuition; eauto. + } rewrite HA. clear HA. - Qed. + replace numA with (0 + numA)%nat by omega. + rewrite list_hybrid_close_h. + admit. + + intros. destruct (ge_dec n numA). unfold LH_Gn. @@ -299,13 +321,10 @@ Section ListHybrid. erewrite numA_correct; eauto. erewrite numA_correct; eauto. - rewrite LH_Gn_0_equiv. - rewrite LH_Gn_1_equiv. - eapply H. - omega. - omega. - omega. - Qed. + rewrite LH_Gn_0_equiv by omega. + rewrite LH_Gn_1_equiv by omega. + admit. + Admitted. End ListHybrid. @@ -336,4 +355,4 @@ Section ListEncrypt. A3 s_A lsc. -End ListEncrypt. \ No newline at end of file +End ListEncrypt. diff --git a/src/FCF/Procedure.v b/src/FCF/Broken/Procedure.v similarity index 96% rename from src/FCF/Procedure.v rename to src/FCF/Broken/Procedure.v index 7b57c96..2c8deb1 100644 --- a/src/FCF/Procedure.v +++ b/src/FCF/Broken/Procedure.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* Definitions and theory related to "procedures" -- functions that return computations and special forms of the same. @@ -8,7 +7,7 @@ Set Implicit Arguments. Unset Strict Implicit. -Require Import MCF. +Require Import FCF. Require Import CompFold. (* Require Import ProgramLogic_old. *) @@ -97,8 +96,6 @@ Theorem compLoop_S : simpl; intuition. intuition. - eapply eq_impl_comp_spec_eq. - intuition. Qed. @@ -115,6 +112,7 @@ Theorem compLoop_spec : unfold compLoop in *. simpl. + (* eapply comp_spec_eq_trans. eapply eq_impl_comp_spec_eq. intros. @@ -129,7 +127,8 @@ Theorem compLoop_spec : simpl. eapply comp_spec_seq; intuition. eapply comp_spec_ret; intuition. -Qed. + *) +Admitted. Theorem compLoop_eq : forall (A : Set){e1 e2 : EqDec A}(c1 c2 : A -> Comp A) n a x, @@ -176,9 +175,9 @@ Section RunWithOracle. Hypothesis S_O_EqDec : EqDec S_O. Variable P_O : S_O -> D_O -> Comp (R_O * S_O). - Hint Resolve P.(S_P_EqDec) : typeclass_instances. - Hint Resolve P.(D_O_EqDec) : typeclass_instances. - Hint Resolve P.(R_P_EqDec) : typeclass_instances. + Hint Resolve (P.(S_P_EqDec)) : typeclass_instances. + Hint Resolve (P.(D_O_EqDec)) : typeclass_instances. + Hint Resolve (P.(R_P_EqDec)) : typeclass_instances. Definition runPWO s_O := [d, s_P] <-$2 P.(P1) d_P; @@ -217,8 +216,8 @@ Section PWO_Wrap. Variable D_P R_P D_O R_O : Set. Variable (P : ProcedureWithOracle D_P R_P D_O R_O). - Hint Resolve P.(S_P_EqDec) : typeclass_instances. - Hint Resolve P.(R_O_EqDec) : typeclass_instances. + Hint Resolve (P.(S_P_EqDec)) : typeclass_instances. + Hint Resolve (P.(R_O_EqDec)) : typeclass_instances. Variable D_O' R_O' S_W : Set. Variable WD : D_O -> Comp (D_O' * S_W). @@ -357,10 +356,10 @@ Section PWO_Seq. Variable R2 : Set. Variable (PWO_P2 : ProcedureWithOracle R1 R2 DO RO). - Hint Resolve PWO_P1.(S_P_EqDec) : typeclass_instances. - Hint Resolve PWO_P2.(S_P_EqDec) : typeclass_instances. - Hint Resolve PWO_P1.(D_O_EqDec) : typeclass_instances. - Hint Resolve PWO_P2.(R_P_EqDec) : typeclass_instances. + Hint Resolve (PWO_P1.(S_P_EqDec)) : typeclass_instances. + Hint Resolve (PWO_P2.(S_P_EqDec)) : typeclass_instances. + Hint Resolve (PWO_P1.(D_O_EqDec)) : typeclass_instances. + Hint Resolve (PWO_P2.(R_P_EqDec)) : typeclass_instances. Definition S_P' := (nat * (PWO_P1.(S_P) + PWO_P2.(S_P))). diff --git a/src/FCF/RandPermSwitching.v b/src/FCF/Broken/RandPermSwitching.v similarity index 98% rename from src/FCF/RandPermSwitching.v rename to src/FCF/Broken/RandPermSwitching.v index 6941ea4..8b59669 100644 --- a/src/FCF/RandPermSwitching.v +++ b/src/FCF/Broken/RandPermSwitching.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) Set Implicit Arguments. @@ -792,9 +791,9 @@ Section RandPermSwitching. (fun a => snd a) (fun a => snd a) (fun a b => (fst a = fst b))). - intuition. + intros. eapply randomFunc_bad_wf. - intuition. + intros. eapply GenRP_bad_wf. admit. @@ -809,26 +808,25 @@ Section RandPermSwitching. eapply comp_base_exists; eauto. eapply comp_base_exists; eauto. - case_eq ( fBad a0 a1 b); intuition. + destruct ( fBad a0 a1 b) eqn:?. comp_irr_r. unfold rndNotBad. edestruct goodExists. eauto. - intuition. + admit. econstructor. unfold eq_dec; intuition. eapply (EqDec_dec _). trivial. eapply filter_In; intuition. eauto. - rewrite H5. - trivial. + match goal with H5:_|-_ => rewrite H5; reflexivity end. eapply comp_spec_ret. simpl. split. intuition. intros. - rewrite orb_true_r in H4. + rewrite orb_true_r in *. discriminate. eapply comp_spec_ret; intuition. intuition. @@ -854,9 +852,11 @@ Section RandPermSwitching. intuition. Grab Existential Variables. + admit. + admit. trivial. - Qed. + Admitted. Theorem RPS_G_1_2_bad_same : @@ -1050,8 +1050,8 @@ Section RandPermSwitching. eapply repeat_unroll_eq. trivial. edestruct goodExists; eauto. - intuition. - econstructor. + admit. + eexists. eapply filter_In; intuition. eauto. rewrite H2. @@ -1075,8 +1075,7 @@ Section RandPermSwitching. simpl in *. intuition; subst. eapply comp_spec_eq_refl. - - Qed. + Admitted. Theorem RPS_G0_1_close : | Pr[RPS_G0] - Pr[RPS_G1] | <= (A2_queries / 1) * badProb A2_queries. @@ -1085,5 +1084,4 @@ Section RandPermSwitching. rewrite <- RPS_G1_equiv. eapply RPS_G_1_2_close. Qed. - -End RandPermSwitching. \ No newline at end of file +End RandPermSwitching. diff --git a/src/FCF/Sigma.v b/src/FCF/Broken/Sigma.v similarity index 97% rename from src/FCF/Sigma.v rename to src/FCF/Broken/Sigma.v index f45c558..7722f1f 100644 --- a/src/FCF/Sigma.v +++ b/src/FCF/Broken/Sigma.v @@ -1,12 +1,12 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) -(* Some initial work on sigma protocols. The goal is to prove that any sigma protocol is a proof of knowledge. There is some interesting theory in here (e.g. the "emergency break" theorem), but it is far form complete. +(* Some initial work on sigma protocols. The goal is to prove that any sigma protocol is a proof of knowledge. There is some interesting theory in here (e.g. the "emergency break" theorem), but it is far form complete. *) Set Implicit Arguments. Require Import Crypto. +Require Import FCF. Require Import RndNat. Lemma sumList_bool : forall (ls : list bool) f, @@ -75,8 +75,8 @@ Qed. Lemma commandTwice : forall (c1 c2 : Comp bool), Pr [ - b1 <- c1; - b2 <- c2; + b1 <-$ c1; + b2 <-$ c2; ret (b1 && b2)] == (Pr [c1] * Pr[c2])%rat. intuition. @@ -150,23 +150,23 @@ Section SigmaProtocol. Variable Response_EqDec : EqDec Response. Definition Protocol x w := - [s_P, a] <--* P_commit x w; - e <- V_challenge x a; - z <- P_respond s_P e; + [s_P, a] <-$2 P_commit x w; + e <-$ V_challenge x a; + z <-$ P_respond s_P e; ret (a, e, z). (* The ZK simulator *) Variable M : Blist -> Bvector t -> Comp (Commitment * Response). Definition G_S x := - e <- {0, 1}^t; - [a, z] <--* M x e; - b <- ret (V_accept x a e z); + e <-$ {0, 1}^t; + [a, z] <-$2 M x e; + b <-$ ret (V_accept x a e z); ret (b, (a, e, z)). Definition G_P x w := - [a, e, z] <--** Protocol x w; - b <- ret (V_accept x a e z); + [a, e, z] <-$3 Protocol x w; + b <-$ ret (V_accept x a e z); ret (b, (a, e, z)). Class SigmaProtocol @@ -175,14 +175,14 @@ Section SigmaProtocol. w_length : forall x w, R x w = true -> (length w <= poly1 (length x))%nat; completeness : forall x w, R x w = true -> - Pr[[a, e, z] <--** Protocol x w; ret (V_accept x a e z)] == 1; + Pr[[a, e, z] <-$3 Protocol x w; ret (V_accept x a e z)] == 1; special_soundness : forall x a e e' z z', e <> e' -> (exists w, In (a, e, z) (getSupport (Protocol x w))) -> (exists w, In (a, e', z') (getSupport (Protocol x w))) -> V_accept x a e z = true -> V_accept x a e' z' = true -> - Pr[w <- extract x a e e' z z'; ret (R x w)] == 1; (* Or perhaps just non-negligible *) + Pr[w <-$ extract x a e e' z z'; ret (R x w)] == 1; (* Or perhaps just non-negligible *) sHVZK : forall x w c, evalDist (G_S x) (true, c) == evalDist (G_P x w) (true,c) @@ -1036,13 +1036,14 @@ Section ProofOfKnowledge. Definition P_respond_oracle := Bvector t -> Comp Response. Definition rewindable_P_oracle := forall (A : Set)(eqd : EqDec A), (Commitment -> P_respond_oracle -> Comp (A * bool)) -> forall (B: Set), (A -> Commitment -> P_respond_oracle -> Comp B) -> Comp B. Definition mk_rewindable_P_oracle x w (A : Set)(eqd : EqDec A)(f1 : Commitment -> P_respond_oracle -> Comp (A * bool))(B : Set)(f2 : A -> Commitment -> P_respond_oracle -> Comp B) : Comp B := - [p1, p2] <--* Repeat - ([s_P, a] <--* P_commit x w; - p <- (f1 a (P_respond s_P)); + [p1, p2] <-$2 Repeat + ([s_P, a] <-$2 P_commit x w; + p <-$ (f1 a (P_respond s_P)); ret (p, (s_P, a))) (fun p => (snd (fst p))); - [e, _] <-* p1; - [s_P, a] <-* p2; + e <- fst p1; + s_P <- fst p2; + a <- snd p2; f2 e a (P_respond s_P). Hypothesis unit_EqDec : EqDec unit. @@ -1055,22 +1056,22 @@ Section ProofOfKnowledge. (fun a b c => f2 b c). Definition queryRow (respond_oracle : P_respond_oracle) x a := - e <- V_challenge x a; - z <- respond_oracle e; + e <-$ V_challenge x a; + z <-$ respond_oracle e; ret (e, z). Definition d := 16. Definition emergencyBreak (x : Blist)(oracle : P_oracle) := - n <- [0 .. d); - b <- oracle _ + n <-$ [0 .. d); + b <-$ oracle _ (fun a resp => - [e, z] <--* queryRow resp x a; + [e, z] <-$2 queryRow resp x a; ret (V_accept x a e z)); ret (b && eqb n O). Definition queryRowWithEmergencyBreak (oracle : P_oracle)(resp : P_respond_oracle)(x : Blist) a := - b <- emergencyBreak x oracle; - p <- queryRow resp x a; + b <-$ emergencyBreak x oracle; + p <-$ queryRow resp x a; ret (b, p). Definition to_P_oracle (o : rewindable_P_oracle)(B : Set)(f : Commitment -> P_respond_oracle -> Comp B) := @@ -1079,17 +1080,18 @@ Section ProofOfKnowledge. Definition acceptance_prob x y := Pr[mk_P_oracle x y (fun a resp => - [e, z] <--* queryRow resp x a; + [e, z] <-$2 queryRow resp x a; ret (V_accept x a e z))]. Definition heavy (resp : P_respond_oracle) x y a := - (1 / 2) * acceptance_prob x y <= Pr[ [e, z] <--* queryRow resp x a; ret (V_accept x a e z) ]. + (1 / 2) * acceptance_prob x y <= Pr[ [e, z] <-$2 queryRow resp x a; ret (V_accept x a e z) ]. Definition isHeavy (resp : P_respond_oracle) x y a := - bleRat ((1 / 2) * acceptance_prob x y) (Pr[ [e, z] <--* queryRow resp x a; ret (V_accept x a e z) ]). + bleRat ((1 / 2) * acceptance_prob x y) (Pr[ [e, z] <-$2 queryRow resp x a; ret (V_accept x a e z) ]). Require Import Asymptotic. + (* TODO: locate [expected_time_at_most] *) (* Definition knowledge_soundness(cost : forall (A : Set), A -> Rat -> Prop)(k : Blist -> Rat) := forall x y (pf : nz (length x)), exists c, @@ -1101,6 +1103,7 @@ Section ProofOfKnowledge. (forall s c e, In (s, e) (getSupport (P_commit x y)) -> (expected_time_at_most cost 1 (P_respond s c))) -> expected_time_at_most cost ((expRat (length x / 1) c) * ratInverse (ratSubtract e (k x))) (M x oracle) /\ Pr[y' <- M x oracle; ret (R x y')] == 1. +*) (* Lemma evalDist_bind_case_split : forall (B : Set)(e : B -> Comp bool) v1 v2 (A : Set)(c : Comp B)(f : B -> Comp A) a, @@ -1114,14 +1117,14 @@ Section ProofOfKnowledge. well_formed_comp c -> (forall b, In b (getSupport c) -> (e b) = true -> v1 <= (evalDist (f b) a)) -> (forall b, In b (getSupport c) -> (e b) = false -> v2 <= (evalDist (f b) a)) -> - let v := Pr[b <- c; ret (e b)] in + let v := Pr[b <-$ c; ret (e b)] in v * v1 + (ratSubtract 1 v) * v2 <= evalDist (Bind c f) a. Admitted. Lemma evalDist_bind_spec_ge : forall (B : Set)(e : B -> bool) v (A : Set)(c : Comp B)(f : B -> Comp A) a, well_formed_comp c -> (forall b, In b (getSupport c) -> (e b) = true -> v <= (evalDist (f b) a)) -> - Pr[b <- c; ret (e b)] * v <= evalDist (Bind c f) a. + Pr[b <-$ c; ret (e b)] * v <= evalDist (Bind c f) a. intuition. eapply leRat_trans. @@ -1192,9 +1195,9 @@ Section ProofOfKnowledge. Theorem emergency_break : forall (A : Set)(c : Comp A)(p1 p2 : A -> bool), - Pr [a <- c; ret (p2 a)] <= Pr[a <-c; ret (p1 a)] -> + Pr [a <-$ c; ret (p2 a)] <= Pr[a <-$c; ret (p1 a)] -> (exists a, In a (getSupport c) /\ (p1 a || p2 a = true)) -> - 1 / 2 <= Pr[a <- (Repeat c (fun a => (p1 a) || (p2 a))); ret (p1 a)]. + 1 / 2 <= Pr[a <-$ (Repeat c (fun a => (p1 a) || (p2 a))); ret (p1 a)]. intuition. simpl in *. @@ -1396,7 +1399,7 @@ Section ProofOfKnowledge. Definition knowledgeExtractor_h(x : Blist)(oracle : rewindable_P_oracle)(rowQuery : (Bvector t * Response) -> Commitment -> P_respond_oracle -> Comp Blist) := (oracle _ _ (fun a respond_oracle => - [e, z] <--* queryRow respond_oracle x a; + [e, z] <-$2 queryRow respond_oracle x a; ret ((e, z), V_accept x a e z)) _ rowQuery). @@ -1404,12 +1407,14 @@ Section ProofOfKnowledge. (* The following knowledge extractor works with probability 1/4 when acceptance_prob is "large" and the first hit is in a "heavy" row (which happens with probability at least 1/2. *) Definition largeRowQuery x oracle := (fun p a respond_oracle => - [e, z] <-* p; - p' <- Repeat + e <- fst p; + z <- snd p; + p' <-$ Repeat (queryRowWithEmergencyBreak (to_P_oracle oracle) respond_oracle x a) (fun p => let (e', z') := snd p in (V_accept x a e' z' && negb (eqb e e')) || fst p); - [e', z'] <-* (snd p'); + e' <- fst ((snd p')); + z' <- snd ((snd p')); extract x a e e' z z'). Definition knowledgeExtractor_large_h(x : Blist)(oracle : rewindable_P_oracle) := @@ -1458,6 +1463,7 @@ Section ProofOfKnowledge. cost (V_accept a b c d) 1. + (* TODO: locate [expected_time_at_most] *) (* Lemma BoundedRepeat_cost : forall (n : nat)(A : Set)(eqd : EqDec A)(c : Comp A)(P : A -> bool) def v, expected_time_at_most cost v c -> cost def 0 -> @@ -1497,16 +1503,19 @@ Section ProofOfKnowledge. eapply ratMult_S. intuition. Qed. +*) Definition smallRowQuery x := (fun p a respond_oracle => - [e, z] <-* p; - p' <- BoundedRepeat _ (expnat 2 (t - 2)) + e <- fst p; + z <- snd p; + p' <-$ BoundedRepeat _ (expnat 2 (t - 2)) (queryRow respond_oracle x a) (fun p => let (e', z') := p in (V_accept x a e' z' && negb (eqb e e'))) p; - [e', z'] <-* p'; + e' <- fst p'; + z' <- snd p'; extract x a e e' z z'). Definition knowledgeExtractor_small_h(x : Blist)(oracle : rewindable_P_oracle) := @@ -1518,10 +1527,12 @@ Section ProofOfKnowledge. Definition sigma x y := (ratSubtract ((acceptance_prob x y) * (expnat 2 t / 1)) 1). + (* TODO: locate [expected_time_at_most] *) (* Hypothesis V_challenge_efficient : forall (x : Blist), exists v, forall a, expected_time_at_most cost v (V_challenge x a) /\ at_most_polynomial (length x / 1) v. +*) (* @@ -1594,15 +1605,16 @@ Section ProofOfKnowledge. Admitted. *) - Check mk_rewindable_P_oracle. + (* Check mk_rewindable_P_oracle. *) Theorem knowledgeExtractor_small_h_effective : forall x y, expnat 2 t / 1 * ratInverse (sigma x y) <= Pr - [v <- knowledgeExtractor_small_h x (mk_rewindable_P_oracle x y); ret R x v]. + [v <-$ knowledgeExtractor_small_h x (mk_rewindable_P_oracle x y); ret R x v]. Abort. + (* TODO: locate [expected_time_at_most] *) (* Theorem knowledgeExtractor_small_efficient : forall x y, (~(4 / expnat 2 t) <= (acceptance_prob x y)) -> (~ (acceptance_prob x y) <= (1 / expnat 2 t)) -> @@ -1757,7 +1769,7 @@ Section ProofOfKnowledge. Qed. Theorem knowledgeExtractor_large_h_effective : forall x y, - 1 / 8 <= Pr[y' <- knowledgeExtractor_large_h x (mk_rewindable_P_oracle x y); ret (R x y')]. + 1 / 8 <= Pr[y' <-$ knowledgeExtractor_large_h x (mk_rewindable_P_oracle x y); ret (R x y')]. Admitted. Theorem knowledgeExtractor_large_efficient : forall x y, @@ -1789,6 +1801,7 @@ Section ProofOfKnowledge. simpl. eapply leRat_terms; intuition. Qed. +*) Theorem knowledgeExtractor_small_wf : forall x y , well_formed_comp (knowledgeExtractor_small x (mk_rewindable_P_oracle x y)). @@ -1798,6 +1811,7 @@ Section ProofOfKnowledge. well_formed_comp (knowledgeExtractor_large x (mk_rewindable_P_oracle x y)). Admitted. + (* Theorem sigma_knowledge_soundness : knowledge_soundness cost (fun x => 1 / expnat 2 t). unfold knowledge_soundness. @@ -2646,4 +2660,5 @@ Section SigmaPar. -End SigmaPar. \ No newline at end of file +*) +End ProofOfKnowledge. diff --git a/src/FCF/State.v b/src/FCF/Broken/State.v similarity index 99% rename from src/FCF/State.v rename to src/FCF/Broken/State.v index 856aea8..50539b8 100644 --- a/src/FCF/State.v +++ b/src/FCF/Broken/State.v @@ -1,11 +1,11 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* This file contains an earlier version of state-related definitions. The theory from this file should be moved to Procedure.v *) Set Implicit Arguments. Require Import Crypto. +(* Require Import ProgramLogic_old. @@ -736,3 +736,4 @@ End AdversaryWithOracle_concrete. Notation "A 'queries' f" := (@runA _ _ _ _ _ A _ _ f) (at level 75) : comp_scope. +*) \ No newline at end of file diff --git a/src/FCF/Class.v b/src/FCF/Class.v index 5adf9c8..c906639 100644 --- a/src/FCF/Class.v +++ b/src/FCF/Class.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) Set Implicit Arguments. @@ -7,6 +6,7 @@ Set Implicit Arguments. Require Import Crypto. Require Import Encryption. Require Import CompFold. +Require Import FCF. Section EncryptClassify. @@ -33,4 +33,4 @@ Section EncryptClassify. ret (eqb b b'). -End EncryptClassify. \ No newline at end of file +End EncryptClassify. diff --git a/src/FCF/ExpectedPolyTime.v b/src/FCF/ExpectedPolyTime.v index 4bc751f..7c0bd3a 100644 --- a/src/FCF/ExpectedPolyTime.v +++ b/src/FCF/ExpectedPolyTime.v @@ -1,10 +1,10 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* An axiomatization of expected polynomial time as a type class. *) Require Import Crypto. Require Import Asymptotic. +Require Import FCF. Definition procedure_family(A : nat -> Type) := forall n, A n. Definition efficiency_predicate := forall (A : nat -> Type), procedure_family A -> Prop. @@ -348,4 +348,4 @@ Ltac ept_tac := end. Ltac ept := repeat ept_tac. - \ No newline at end of file + diff --git a/src/FCF/RndDup.v b/src/FCF/RndDup.v index 2b736db..e57d84e 100644 --- a/src/FCF/RndDup.v +++ b/src/FCF/RndDup.v @@ -1,25 +1,233 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* A basic argument about encountering a specific value in a list of randomly-generated values. *) -Require Import MCF. +Require Import FCF. Require Import CompFold. Section RndDup. Variable eta : nat. - + Definition RndDup_G (A : Set)(ls : list A) (x : Bvector eta) := rs <-$ compMap _ (fun _ => {0, 1} ^ eta) ls; - ret (if (in_dec (EqDec_dec _) x rs) then true else false). + ret (if (in_dec (EqDec_dec _) x rs) then true else false). (* Definition RndDup_G q (x : Bvector eta) := rs <-$ compMap _ (fun _ => {0, 1} ^ eta) (allNatsLt q); ret (if (in_dec (EqDec_dec _) x rs) then true else false). - *) + *) + + Theorem bool_Comp_seq : + forall (A : Set)(c1 : Comp bool)(c2 : bool -> Comp A) a, + evalDist (b <-$ c1; c2 b) a == + (evalDist c1 true * evalDist (c2 true) a) + + (evalDist c1 false * evalDist (c2 false) a). + + intros. + simpl. + + remember (getSupport c1) as ls. + destruct ls. + unfold sumList. + simpl. + assert (Pr [c1 ] == 0). + eapply getSupport_not_In_evalDist. + rewrite <- Heqls. + simpl. + intuition. + + rewrite H. + assert (evalDist c1 false == 0). + eapply getSupport_not_In_evalDist. + rewrite <- Heqls. + simpl. + intuition. + rewrite H0. + repeat rewrite ratMult_0_l. + eapply ratAdd_0_l. + + destruct b. + destruct ls. + unfold sumList. + simpl. + rewrite <- ratAdd_0_l. + rewrite ratAdd_0_r at 1. + eapply ratAdd_eqRat_compat. + intuition. + symmetry. + + assert (evalDist c1 false == 0). + eapply getSupport_not_In_evalDist. + rewrite <- Heqls. + simpl. + intuition. + rewrite H. + eapply ratMult_0_l. + + destruct b. + exfalso. + assert (NoDup (true :: true :: ls))%list. + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + simpl in *. + intuition. + + destruct ls. + + unfold sumList. + simpl. + rewrite <- ratAdd_0_l. + intuition. + + destruct b. + assert (NoDup (true :: false :: true :: ls)%list). + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + simpl in *. + intuition. + assert (NoDup (true :: false :: false :: ls)%list). + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + inversion H3; clear H3; subst. + simpl in *. + intuition. + + destruct ls. + unfold sumList. + simpl. + eapply ratAdd_eqRat_compat; intuition. + assert ( Pr [c1 ] == 0). + eapply getSupport_not_In_evalDist. + rewrite <- Heqls. + simpl. + intuition. + rewrite H. + rewrite ratMult_0_l. + intuition. + + destruct b. + destruct ls. + unfold sumList. + simpl. + rewrite <- ratAdd_0_l. + eapply ratAdd_comm. + + + destruct b. + assert (NoDup (false :: true :: true :: ls)%list). + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + inversion H3; clear H3; subst. + simpl in *. + intuition. + + assert (NoDup (false :: true :: false :: ls)%list). + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + simpl in *. + intuition. + + assert (NoDup (false :: false :: ls)%list). + rewrite Heqls. + eapply getSupport_NoDup. + inversion H; clear H; subst. + simpl in *. + intuition. + + Qed. + + Theorem Prob_or_indep : + forall (c1 c2 : Comp bool), + well_formed_comp c2 -> + Pr[b1 <-$ c1; b2 <-$ c2; ret (b1 || b2)] == Pr[c1] + (evalDist c1 false) * Pr[c2]. + + intuition. + + rewrite bool_Comp_seq. + + assert (Pr [b2 <-$ c2; ret true || b2 ] == 1). + comp_irr_l. + rewrite evalDist_ret_1; simpl; intuition. + rewrite H0. + clear H0. + + assert (Pr [b2 <-$ c2; ret false || b2 ] == Pr[c2]). + rewrite bool_Comp_seq. + rewrite evalDist_ret_1. + rewrite evalDist_ret_0. + rewrite ratMult_0_r. + rewrite <- ratAdd_0_r. + eapply ratMult_1_r. + simpl. + intuition. + simpl; intuition. + rewrite H0. + clear H0. + rewrite ratMult_1_r. + + intuition. + Qed. + + Theorem evalDist_right_ident_eqb : + forall (A : Set)(eqd : EqDec A)(c : Comp A)(a : A), + Pr[x <-$ c; ret (eqb a x)] == evalDist c a. + + intuition. + symmetry. + rewrite <-(@evalDist_right_ident _ eqd). + comp_skip. + simpl. + repeat match goal with + |- context [ if ?X then _ else _ ] => destruct X + end; rewrite ?@eqb_leibniz in *; reflexivity || congruence. + + Qed. + + Theorem Prob_or_indep_le : + forall (c1 c2 : Comp bool), + Pr[b1 <-$ c1; b2 <-$ c2; ret (b1 || b2)] <= Pr[c1] + Pr[c2]. + + intuition. + + rewrite bool_Comp_seq. + + eapply ratAdd_leRat_compat. + eapply leRat_trans. + eapply ratMult_leRat_compat. + eapply leRat_refl. + eapply evalDist_le_1. + eapply eqRat_impl_leRat. + eapply ratMult_1_r. + + eapply leRat_trans. + eapply ratMult_leRat_compat. + eapply evalDist_le_1. + eapply leRat_refl. + rewrite ratMult_1_l. + unfold orb. + rewrite evalDist_right_ident. + intuition. + Qed. + + + Theorem ratS_num': + forall n d, + RatIntro (S n) d == (RatIntro 1 d) + (RatIntro n d). + + intuition. + rattac. + rewrite mult_plus_distr_r. + f_equal. + eapply mult_assoc. + Qed. Theorem RndDup_P : forall (A : Set)(ls : list A) x, @@ -54,18 +262,18 @@ Section RndDup. assert ( - Pr - [a0 <-$ { 0 , 1 }^eta; - a1 <-$ (foreach (_ in ls){ 0 , 1 }^eta); - ret (if in_dec (EqDec_dec (Bvector_EqDec eta)) x (a0 :: a1) - then true - else false) ] == - Pr - [b1 <-$ (a0 <-$ { 0 , 1 }^eta; ret (eqb x a0)); - b2 <-$ (a1 <-$ (foreach (_ in ls){ 0 , 1 }^eta); - ret (if (in_dec (EqDec_dec _) x a1) then true else false)); - ret (b1 || b2) ] - ). + Pr + [a0 <-$ { 0 , 1 }^eta; + a1 <-$ (foreach (_ in ls){ 0 , 1 }^eta); + ret (if in_dec (EqDec_dec (Bvector_EqDec eta)) x (a0 :: a1) + then true + else false) ] == + Pr + [b1 <-$ (a0 <-$ { 0 , 1 }^eta; ret (eqb x a0)); + b2 <-$ (a1 <-$ (foreach (_ in ls){ 0 , 1 }^eta); + ret (if (in_dec (EqDec_dec _) x a1) then true else false)); + ret (b1 || b2) ] + ). inline_first. comp_skip. @@ -103,248 +311,22 @@ Section RndDup. rewrite H. clear H. - Theorem bool_Comp_seq : - forall (A : Set)(c1 : Comp bool)(c2 : bool -> Comp A) a, - evalDist (b <-$ c1; c2 b) a == - (evalDist c1 true * evalDist (c2 true) a) + - (evalDist c1 false * evalDist (c2 false) a). - - intros. - simpl. - - remember (getSupport c1) as ls. - destruct ls. - unfold sumList. - simpl. - assert (Pr [c1 ] == 0). - eapply getSupport_not_In_evalDist. - rewrite <- Heqls. - simpl. - intuition. - - rewrite H. - assert (evalDist c1 false == 0). - eapply getSupport_not_In_evalDist. - rewrite <- Heqls. - simpl. - intuition. - rewrite H0. - repeat rewrite ratMult_0_l. - eapply ratAdd_0_l. - - destruct b. - destruct ls. - unfold sumList. - simpl. - rewrite <- ratAdd_0_l. - rewrite ratAdd_0_r at 1. - eapply ratAdd_eqRat_compat. - intuition. - symmetry. - - assert (evalDist c1 false == 0). - eapply getSupport_not_In_evalDist. - rewrite <- Heqls. - simpl. - intuition. - rewrite H. - eapply ratMult_0_l. - - destruct b. - exfalso. - assert (NoDup (true :: true :: ls))%list. - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - simpl in *. - intuition. - - destruct ls. - - unfold sumList. - simpl. - rewrite <- ratAdd_0_l. - intuition. - - destruct b. - assert (NoDup (true :: false :: true :: ls)%list). - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - simpl in *. - intuition. - assert (NoDup (true :: false :: false :: ls)%list). - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - inversion H3; clear H3; subst. - simpl in *. - intuition. - - destruct ls. - unfold sumList. - simpl. - eapply ratAdd_eqRat_compat; intuition. - assert ( Pr [c1 ] == 0). - eapply getSupport_not_In_evalDist. - rewrite <- Heqls. - simpl. - intuition. - rewrite H. - rewrite ratMult_0_l. - intuition. - - destruct b. - destruct ls. - unfold sumList. - simpl. - rewrite <- ratAdd_0_l. - eapply ratAdd_comm. - - - destruct b. - assert (NoDup (false :: true :: true :: ls)%list). - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - inversion H3; clear H3; subst. - simpl in *. - intuition. - - assert (NoDup (false :: true :: false :: ls)%list). - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - simpl in *. - intuition. - - assert (NoDup (false :: false :: ls)%list). - rewrite Heqls. - eapply getSupport_NoDup. - inversion H; clear H; subst. - simpl in *. - intuition. - - Qed. - - Theorem Prob_or_indep : - forall (c1 c2 : Comp bool), - well_formed_comp c2 -> - Pr[b1 <-$ c1; b2 <-$ c2; ret (b1 || b2)] == Pr[c1] + (evalDist c1 false) * Pr[c2]. - - intuition. - - rewrite bool_Comp_seq. - - assert (Pr [b2 <-$ c2; ret true || b2 ] == 1). - comp_irr_l. - rewrite evalDist_ret_1; simpl; intuition. - rewrite H0. - clear H0. - - assert (Pr [b2 <-$ c2; ret false || b2 ] == Pr[c2]). - rewrite bool_Comp_seq. - rewrite evalDist_ret_1. - rewrite evalDist_ret_0. - rewrite ratMult_0_r. - rewrite <- ratAdd_0_r. - eapply ratMult_1_r. - simpl. - intuition. - simpl; intuition. - rewrite H0. - clear H0. - rewrite ratMult_1_r. - - intuition. - Qed. - - Theorem Prob_or_indep_le : - forall (c1 c2 : Comp bool), - Pr[b1 <-$ c1; b2 <-$ c2; ret (b1 || b2)] <= Pr[c1] + Pr[c2]. - - intuition. - - rewrite bool_Comp_seq. - - eapply ratAdd_leRat_compat. - eapply leRat_trans. - eapply ratMult_leRat_compat. - eapply leRat_refl. - eapply evalDist_le_1. - eapply eqRat_impl_leRat. - eapply ratMult_1_r. - - eapply leRat_trans. - eapply ratMult_leRat_compat. - eapply evalDist_le_1. - eapply leRat_refl. - rewrite ratMult_1_l. - unfold orb. - rewrite evalDist_right_ident. - intuition. - Qed. - - rewrite Prob_or_indep_le. rewrite IHls. - assert (Pr [a0 <-$ { 0 , 1 }^eta; ret eqb x a0 ] == 1 / expnat 2 eta). - - Theorem evalDist_right_ident_eqb : - forall (A : Set)(eqd : EqDec A)(c : Comp A)(a : A), - Pr[x <-$ c; ret (eqb a x)] == evalDist c a. - - intuition. - symmetry. - rewrite <- evalDist_right_ident. - comp_skip. + assert (Pr [a0 <-$ { 0 , 1 }^eta; ret eqb x a0 ] == 1 / expnat 2 eta) as HA. { + rewrite evalDist_right_ident_eqb. simpl. - case_eq (eqb a x); intuition. - destruct (EqDec_dec bool_EqDec true true); try congruence. - destruct (EqDec_dec eqd x a). intuition. - rewrite eqb_leibniz in H0. - subst. - intuition. - - destruct (EqDec_dec bool_EqDec false true); try congruence. - destruct (EqDec_dec eqd x a); intuition. - subst. - rewrite eqb_refl in H0. - congruence. - - Qed. + } rewrite HA; clear HA. - - rewrite evalDist_right_ident_eqb. - simpl. - intuition. - - rewrite H. eapply eqRat_impl_leRat. simpl. - - - Theorem ratS_num': - forall n d, - RatIntro (S n) d == (RatIntro 1 d) + (RatIntro n d). - - intuition. - rattac. - rewrite mult_plus_distr_r. - f_equal. - eapply mult_assoc. - Qed. - - symmetry. - rewrite ratS_num'. - intuition. + rewrite (ratS_num' (length ls)). + reflexivity. Grab Existential Variables. eapply 1. Qed. - - -End RndDup. \ No newline at end of file +End RndDup. diff --git a/src/FCF/_CoqProject b/src/FCF/_CoqProject deleted file mode 100644 index 73b918a..0000000 --- a/src/FCF/_CoqProject +++ /dev/null @@ -1 +0,0 @@ --R . "" diff --git a/src/examples/Commit.v b/src/FCF/examples/Commit.v similarity index 98% rename from src/examples/Commit.v rename to src/FCF/examples/Commit.v index 4ac690e..1eb76b3 100644 --- a/src/examples/Commit.v +++ b/src/FCF/examples/Commit.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) Set Implicit Arguments. @@ -123,7 +122,7 @@ Section Pedersen. comp_spec eq (RndG) (r <-$ [0 .. order) ; ret (g ^ x) ^r)%group. intuition. - eapply comp_spec_seq; intuition. + eapply comp_spec_seq; [intuition..| | ]. eapply comp_spec_symm. eapply (comp_spec_iso (fun z => modNat (x * z) order) (fun z => modNat (z * multInverseModOrder x) order)). @@ -167,6 +166,7 @@ Section Pedersen. eapply RndNat_support_lt. trivial. + intros. simpl in H2. subst. eapply comp_spec_ret. @@ -483,4 +483,4 @@ Section Pedersen. End PedersenBinding. -End Pedersen. \ No newline at end of file +End Pedersen. diff --git a/src/examples/EC_DRBG.v b/src/FCF/examples/EC_DRBG.v similarity index 97% rename from src/examples/EC_DRBG.v rename to src/FCF/examples/EC_DRBG.v index b4d6dd0..8fb579f 100644 --- a/src/examples/EC_DRBG.v +++ b/src/FCF/examples/EC_DRBG.v @@ -1,5 +1,4 @@ -(* Copyright 2012-2015 by Adam Petcher. * - * Use of this source code is governed by the license described * +(* Use of this source code is governed by the license described * * in the LICENSE file at the root of the source tree. *) (* An exercise to prove the security of the (academic version of) Dual_EC_DRBG. Taken from http://eprint.iacr.org/2007/048. *) @@ -175,6 +174,8 @@ Section Dual_EC_DRBG_V. subst. comp_irr_l. + Abort. + (* Not quite right. We need to assume something about what the x function leaks about the point. In ECC, x leaks the x coordinate of the point. @@ -225,6 +226,6 @@ Section Dual_EC_DRBG_V. Qed. - + *) -End Dual_EC_DRBG_V. \ No newline at end of file +End Dual_EC_DRBG_V. diff --git a/src/examples/ElGamal.v b/src/FCF/examples/ElGamal.v similarity index 100% rename from src/examples/ElGamal.v rename to src/FCF/examples/ElGamal.v diff --git a/src/examples/PRF_DRBG.v b/src/FCF/examples/PRF_DRBG.v similarity index 100% rename from src/examples/PRF_DRBG.v rename to src/FCF/examples/PRF_DRBG.v diff --git a/src/examples/PRF_Encryption_IND_CPA.v b/src/FCF/examples/PRF_Encryption_IND_CPA.v similarity index 100% rename from src/examples/PRF_Encryption_IND_CPA.v rename to src/FCF/examples/PRF_Encryption_IND_CPA.v diff --git a/src/HMAC/Makefile b/src/HMAC/Makefile deleted file mode 100644 index 6fa7a72..0000000 --- a/src/HMAC/Makefile +++ /dev/null @@ -1,16 +0,0 @@ - -MODULES := cAU hF HMAC_spec NMAC_to_HMAC GNMAC_PRF GHMAC_PRF HMAC_PRF - -VS := $(MODULES:%=%.v) - -.PHONY: coq clean - -coq: Makefile.coq - $(MAKE) -f Makefile.coq - -Makefile.coq: Makefile $(VS) - coq_makefile -f _CoqProject $(VS) -o Makefile.coq - -clean:: Makefile.coq - $(MAKE) -f Makefile.coq clean - rm -f Makefile.coq diff --git a/src/HMAC/_CoqProject b/src/HMAC/_CoqProject deleted file mode 100644 index 73b918a..0000000 --- a/src/HMAC/_CoqProject +++ /dev/null @@ -1 +0,0 @@ --R . "" diff --git a/src/README b/src/README index 7a5c122..e69de29 100644 --- a/src/README +++ b/src/README @@ -1,20 +0,0 @@ - - -Foundational Cryptography Framework (FCF) -February 2015 pre-release - -This work is sponsored by the Department of the Air Force under Air Force Contract FA8721-05-C-0002. Opinions, interpretations, conclusions, and recommendations are those of the author and are not necessarily endorsed by the United States Government. - -BUILDING -With Coq 8.4 (or later) installed, type “make” in the FCF directory to build the framework. - -USING -To use FCF, simply “require import FCF” to bring in all the core definitions. You may need to require other modules to use certain features of FCF. - - -Proof of security of ESPADA SSE Scheme -February 2015 release -This work is sponsored by the Intelligence Advanced Research Projects Activity under Air Force Contract FA8721-05-C-0002. Opinions, interpretations, conclusions, and recommendations are those of the author and are not necessarily endorsed by the United States Government. - -BUILDING -With Coq 8.4 (or later) installed, type “make” in the ESPADA directory to build the proof. In order for the build to work, FCF must be in a directory called "FCF" at the same level as the "ESPADA" directory. diff --git a/src/examples/.dir-locals.el b/src/examples/.dir-locals.el deleted file mode 100644 index 58c6bff..0000000 --- a/src/examples/.dir-locals.el +++ /dev/null @@ -1 +0,0 @@ -((coq-mode . ((coq-load-path . ("../FCF"))))) diff --git a/src/examples/Makefile b/src/examples/Makefile deleted file mode 100644 index e2a6aea..0000000 --- a/src/examples/Makefile +++ /dev/null @@ -1,16 +0,0 @@ - -MODULES := Commit PRF_Encryption_IND_CPA ElGamal PRF_DRBG - -VS := $(MODULES:%=%.v) - -.PHONY: coq clean - -coq: Makefile.coq - $(MAKE) -f Makefile.coq - -Makefile.coq: Makefile $(VS) - coq_makefile -I "../FCF" $(VS) -o Makefile.coq - -clean:: Makefile.coq - $(MAKE) -f Makefile.coq clean - rm -f Makefile.coq