diff --git a/thys/Multi_Party_Computation/ETP.thy b/thys/Multi_Party_Computation/ETP.thy --- a/thys/Multi_Party_Computation/ETP.thy +++ b/thys/Multi_Party_Computation/ETP.thy @@ -1,100 +1,87 @@ -subsection \ETP definitions\ +subsection \ ETP definitions \ -text \We define Extended Trapdoor Permutations (ETPs) following \cite{DBLP:books/sp/17/Lindell17} and \cite{DBLP:books/cu/Goldreich2004}. -In particular we consider the property of Hard Core Predicates (HCPs).\ +text \ We define Extended Trapdoor Permutations (ETPs) following \cite{DBLP:books/sp/17/Lindell17} and \cite{DBLP:books/cu/Goldreich2004}. +In particular we consider the property of Hard Core Predicates (HCPs). \ theory ETP imports CryptHOL.CryptHOL begin type_synonym ('index,'range) dist2 = "(bool \ 'index \ bool \ bool) \ bool spmf" type_synonym ('index,'range) advP2 = "'index \ bool \ bool \ ('index,'range) dist2 \ 'range \ bool spmf" locale etp = fixes I :: "('index \ 'trap) spmf" \ \samples index and trapdoor\ and domain :: "'index \ 'range set" and range :: "'index \ 'range set" - and f :: "'index \ ('range \ 'range)" \ \permutation\ + and F :: "'index \ ('range \ 'range)" \ \permutation\ + and F\<^sub>i\<^sub>n\<^sub>v :: "'index \ 'trap \ 'range \ 'range" \ \must be efficiently computable\ and B :: "'index \ 'range \ bool" \ \hard core predicate\ assumes dom_eq_ran: "y \ set_spmf I \ domain (fst y) = range (fst y)" and finite_range: "y \ set_spmf I \ finite (range (fst y))" and non_empty_range: "y \ set_spmf I \ range (fst y) \ {}" - and bij_betw: "y \ set_spmf I \ bij_betw (f (fst y)) (domain (fst y)) (range (fst y))" + and bij_betw: "y \ set_spmf I \ bij_betw (F (fst y)) (domain (fst y)) (range (fst y))" and lossless_I: "lossless_spmf I" + and F_f_inv: "y \ set_spmf I \ x \ range (fst y) \ F\<^sub>i\<^sub>n\<^sub>v (fst y) (snd y) (F (fst y) x) = x" begin definition S :: "'index \ 'range spmf" where "S \ = spmf_of_set (range \)" lemma lossless_S: "y \ set_spmf I \ lossless_spmf (S (fst y))" by(simp add: lossless_spmf_def S_def finite_range non_empty_range) lemma set_spmf_S [simp]: "y \ set_spmf I \ set_spmf (S (fst y)) = range (fst y)" by (simp add: S_def finite_range) -definition "f\<^sub>i\<^sub>n\<^sub>v \ = Hilbert_Choice.inv_into (range \) (f \)" +lemma f_inj_on: "y \ set_spmf I \ inj_on (F (fst y)) (range (fst y))" + by(metis bij_betw_def bij_betw dom_eq_ran bij_betw_def bij_betw dom_eq_ran) -lemma f_inj_on: "y \ set_spmf I \ inj_on (f (fst y)) (range (fst y))" - by (metis bij_betw_def bij_betw dom_eq_ran bij_betw_def bij_betw dom_eq_ran) - -lemma range_f: "y \ set_spmf I \ x \ range (fst y) \ f (fst y) x \ range (fst y)" +lemma range_f: "y \ set_spmf I \ x \ range (fst y) \ F (fst y) x \ range (fst y)" by (metis bij_betw bij_betw dom_eq_ran bij_betwE) -lemma f_inv_f [simp]: "y \ set_spmf I \ x \ range (fst y) \ f\<^sub>i\<^sub>n\<^sub>v (fst y) (f (fst y) x) = x" - by (metis bij_betw bij_betw_inv_into_left dom_eq_ran f\<^sub>i\<^sub>n\<^sub>v_def) - -lemma f_inv_f' [simp]: "y \ set_spmf I \ x \ range (fst y) \ Hilbert_Choice.inv_into (range (fst y)) (f (fst y)) (f (fst y) x) = x" - by (metis bij_betw bij_betw_inv_into_left bij_betw dom_eq_ran) - -definition F :: "'index \ 'range \ 'range" - where "F \ x = f \ x" +lemma f_inv_f [simp]: "y \ set_spmf I \ x \ range (fst y) \ F\<^sub>i\<^sub>n\<^sub>v (fst y) (snd y) (F (fst y) x) = x" + by (metis bij_betw bij_betw_inv_into_left dom_eq_ran F_f_inv) -definition F\<^sub>i\<^sub>n\<^sub>v :: "'index \ 'trap \ 'range \ 'range" - where "F\<^sub>i\<^sub>n\<^sub>v \ \ y = Hilbert_Choice.inv_into (range \) (f \) y" - -lemma F_f_inv: "F\<^sub>i\<^sub>n\<^sub>v \ \ y = f\<^sub>i\<^sub>n\<^sub>v \ y" - by(simp add: F\<^sub>i\<^sub>n\<^sub>v_def f\<^sub>i\<^sub>n\<^sub>v_def) - -lemma F_invF [simp]: "x \ set_spmf I \ y \ range (fst x) \ F (fst x) (F\<^sub>i\<^sub>n\<^sub>v (fst x) \ y) = y" - apply(simp add: F_def F\<^sub>i\<^sub>n\<^sub>v_def) - by (metis bij_betw bij_betw_inv_into_right dom_eq_ran) +lemma f_inv_f' [simp]: "y \ set_spmf I \ x \ range (fst y) \ Hilbert_Choice.inv_into (range (fst y)) (F (fst y)) (F (fst y) x) = x" + by (metis bij_betw bij_betw_inv_into_left bij_betw dom_eq_ran) lemma B_F_inv_rewrite: "(B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = (B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = m1)) = m1" by auto lemma uni_set_samp: assumes "y \ set_spmf I" shows "map_spmf (\ x. F (fst y) x) (S (fst y)) = (S (fst y))" (is "?lhs = ?rhs") proof- have rhs: "?rhs = spmf_of_set (range (fst y))" unfolding S_def by(simp) - also have "map_spmf (\ x. f (fst y) x) (spmf_of_set (range (fst y))) = spmf_of_set ((\ x. f (fst y) x) ` (range (fst y)))" + also have "map_spmf (\ x. F (fst y) x) (spmf_of_set (range (fst y))) = spmf_of_set ((\ x. F (fst y) x) ` (range (fst y)))" using f_inj_on assms by (metis map_spmf_of_set_inj_on) - also have "(\ x. f (fst y) x) ` (range (fst y)) = range (fst y)" + also have "(\ x. F (fst y) x) ` (range (fst y)) = range (fst y)" apply(rule endo_inj_surj) using bij_betw by (auto simp add: bij_betw_def dom_eq_ran f_inj_on bij_betw finite_range assms) - finally show ?thesis by(simp add: rhs F_def) + finally show ?thesis by(simp add: rhs) qed text\We define the security property of the hard core predicate (HCP) using a game.\ definition HCP_game :: "('index,'range) advP2 \ bool \ bool \ ('index,'range) dist2 \ bool spmf" where "HCP_game A = (\ \ b\<^sub>\ D. do { (\, \) \ I; x \ S \; b' \ A \ \ b\<^sub>\ D x; let b = B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); return_spmf (b = b')})" definition "HCP_adv A \ b\<^sub>\ D = \((spmf (HCP_game A \ b\<^sub>\ D) True) - 1/2)\" end end diff --git a/thys/Multi_Party_Computation/ETP_OT.thy b/thys/Multi_Party_Computation/ETP_OT.thy --- a/thys/Multi_Party_Computation/ETP_OT.thy +++ b/thys/Multi_Party_Computation/ETP_OT.thy @@ -1,437 +1,426 @@ -subsection \Oblivious transfer constructed from ETPs\ +subsection \ Oblivious transfer constructed from ETPs \ text\Here we construct the OT protocol based on ETPs given in \cite{DBLP:books/sp/17/Lindell17} (Chapter 4) and prove semi honest security for both parties. We show information theoretic security for Party 1 and reduce the security of Party 2 to the HCP assumption.\ theory ETP_OT imports "HOL-Number_Theory.Cong" ETP OT_Functionalities Semi_Honest_Def begin type_synonym 'range viewP1 = "((bool \ bool) \ 'range \ 'range) spmf" type_synonym 'range dist1 = "((bool \ bool) \ 'range \ 'range) \ bool spmf" type_synonym 'index viewP2 = "(bool \ 'index \ (bool \ bool)) spmf" type_synonym 'index dist2 = "(bool \ 'index \ bool \ bool) \ bool spmf" type_synonym ('index, 'range) advP2 = "'index \ bool \ bool \ 'index dist2 \ 'range \ bool spmf" lemma if_False_True: "(if x then False else \ False) \ (if x then False else True)" by simp lemma if_then_True [simp]: "(if b then True else x) \ (\ b \ x)" by simp lemma if_else_True [simp]: "(if b then x else True) \ (b \ x)" by simp lemma inj_on_Not [simp]: "inj_on Not A" by(auto simp add: inj_on_def) -locale ETP_base = - fixes I :: "('index \ 'trap) spmf" - and domain :: "'index \ 'range set" +locale ETP_base = etp: etp I domain range F F\<^sub>i\<^sub>n\<^sub>v B + for I :: "('index \ 'trap) spmf" \ \samples index and trapdoor\ + and domain :: "'index \ 'range set" and range :: "'index \ 'range set" - and f :: "'index \ ('range \ 'range)" - and B :: "'index \ 'range \ bool" - assumes dom_eq_ran: "y \ set_spmf I \ domain (fst y) = range (fst y)" - and finite_range: "y \ set_spmf I \ finite (range (fst y))" - and non_empty_range: "y \ set_spmf I \ range (fst y) \ {}" - and bij_betw: "y \ set_spmf I \ bij_betw (f (fst y)) (domain (fst y)) (range (fst y))" - and lossless_I: "lossless_spmf I" + and B :: "'index \ 'range \ bool" \ \hard core predicate\ + and F :: "'index \ 'range \ 'range" + and F\<^sub>i\<^sub>n\<^sub>v :: "'index \ 'trap \ 'range \ 'range" begin -sublocale etp: etp I domain range f B - unfolding etp_def using dom_eq_ran finite_range non_empty_range bij_betw lossless_I by simp - text\The probabilistic program that defines the protocol.\ definition protocol :: "(bool \ bool) \ bool \ (unit \ bool) spmf" where "protocol input\<^sub>1 \ = do { let (b\<^sub>\, b\<^sub>\') = input\<^sub>1; (\ :: 'index, \ :: 'trap) \ I; x\<^sub>\ :: 'range \ etp.S \; y\<^sub>\' :: 'range \ etp.S \; - let (y\<^sub>\ :: 'range) = etp.F \ x\<^sub>\; - let (x\<^sub>\ :: 'range) = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; - let (x\<^sub>\' :: 'range) = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; + let (y\<^sub>\ :: 'range) = F \ x\<^sub>\; + let (x\<^sub>\ :: 'range) = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; + let (x\<^sub>\' :: 'range) = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; let (\\<^sub>\ :: bool) = xor (B \ x\<^sub>\) b\<^sub>\; let (\\<^sub>\' :: bool) = xor (B \ x\<^sub>\') b\<^sub>\'; return_spmf ((), if \ then xor (B \ x\<^sub>\') \\<^sub>\' else xor (B \ x\<^sub>\) \\<^sub>\)}" lemma correctness: "protocol (m0,m1) c = funct_OT_12 (m0,m1) c" proof- - have "(B \ (etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = (B \ (etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = m1)) = m1" + have "(B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = (B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\') = m1)) = m1" for \ \ y\<^sub>\' by auto then show ?thesis - by(auto simp add: protocol_def funct_OT_12_def Let_def etp.B_F_inv_rewrite bind_spmf_const etp.lossless_S lossless_I lossless_weight_spmfD split_def cong: bind_spmf_cong) + by(auto simp add: protocol_def funct_OT_12_def Let_def etp.B_F_inv_rewrite bind_spmf_const etp.lossless_S local.etp.lossless_I lossless_weight_spmfD split_def cong: bind_spmf_cong) qed -text \Party 1 views\ +text \ Party 1 views \ definition R1 :: "(bool \ bool) \ bool \ 'range viewP1" where "R1 input\<^sub>1 \ = do { let (b\<^sub>0, b\<^sub>1) = input\<^sub>1; (\, \) \ I; x\<^sub>\ \ etp.S \; y\<^sub>\' \ etp.S \; - let y\<^sub>\ = etp.F \ x\<^sub>\; + let y\<^sub>\ = F \ x\<^sub>\; return_spmf ((b\<^sub>0, b\<^sub>1), if \ then y\<^sub>\' else y\<^sub>\, if \ then y\<^sub>\ else y\<^sub>\')}" lemma lossless_R1: "lossless_spmf (R1 msgs \)" - by(simp add: R1_def lossless_I split_def etp.lossless_S Let_def) + by(simp add: R1_def local.etp.lossless_I split_def etp.lossless_S Let_def) definition S1 :: "(bool \ bool) \ unit \ 'range viewP1" where "S1 == (\ input\<^sub>1 (). do { let (b\<^sub>0, b\<^sub>1) = input\<^sub>1; (\, \) \ I; y\<^sub>0 :: 'range \ etp.S \; y\<^sub>1 \ etp.S \; return_spmf ((b\<^sub>0, b\<^sub>1), y\<^sub>0, y\<^sub>1)})" lemma lossless_S1: "lossless_spmf (S1 msgs ())" - by(simp add: S1_def lossless_I split_def etp.lossless_S) + by(simp add: S1_def local.etp.lossless_I split_def etp.lossless_S) -text \Party 2 views\ +text \ Party 2 views \ definition R2 :: "(bool \ bool) \ bool \ 'index viewP2" where "R2 msgs \ = do { let (b0,b1) = msgs; (\, \) \ I; x\<^sub>\ \ etp.S \; y\<^sub>\' \ etp.S \; - let y\<^sub>\ = etp.F \ x\<^sub>\; - let x\<^sub>\ = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; - let x\<^sub>\' = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; + let y\<^sub>\ = F \ x\<^sub>\; + let x\<^sub>\ = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; + let x\<^sub>\' = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; let \\<^sub>\ = (B \ x\<^sub>\) \ (if \ then b1 else b0) ; let \\<^sub>\' = (B \ x\<^sub>\') \ (if \ then b0 else b1); return_spmf (\, \,(\\<^sub>\, \\<^sub>\'))}" lemma lossless_R2: "lossless_spmf (R2 msgs \)" - by(simp add: R2_def split_def lossless_I etp.lossless_S) + by(simp add: R2_def split_def local.etp.lossless_I etp.lossless_S) definition S2 :: "bool \ bool \ 'index viewP2" where "S2 \ b\<^sub>\ = do { (\, \) \ I; x\<^sub>\ \ etp.S \; y\<^sub>\' \ etp.S \; - let x\<^sub>\' = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; + let x\<^sub>\' = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; let \\<^sub>\' = B \ x\<^sub>\'; return_spmf (\, \, (\\<^sub>\, \\<^sub>\'))}" lemma lossless_S2: "lossless_spmf (S2 \ b\<^sub>\)" - by(simp add: S2_def lossless_I etp.lossless_S split_def) + by(simp add: S2_def local.etp.lossless_I etp.lossless_S split_def) -text \Security for Party 1\ +text \ Security for Party 1 \ text\We have information theoretic security for Party 1.\ lemma P1_security: "R1 input\<^sub>1 \ = funct_OT_12 x y \ (\ (s1, s2). S1 input\<^sub>1 s1)" including monad_normalisation proof- have "R1 input\<^sub>1 \ = do { let (b0,b1) = input\<^sub>1; (\, \) \ I; y\<^sub>\' :: 'range \ etp.S \; - y\<^sub>\ \ map_spmf (\ x\<^sub>\. etp.F \ x\<^sub>\) (etp.S \); + y\<^sub>\ \ map_spmf (\ x\<^sub>\. F \ x\<^sub>\) (etp.S \); return_spmf ((b0,b1), if \ then y\<^sub>\' else y\<^sub>\, if \ then y\<^sub>\ else y\<^sub>\')}" by(simp add: bind_map_spmf o_def Let_def R1_def) also have "... = do { let (b0,b1) = input\<^sub>1; (\, \) \ I; y\<^sub>\' :: 'range \ etp.S \; y\<^sub>\ \ etp.S \; return_spmf ((b0,b1), if \ then y\<^sub>\' else y\<^sub>\, if \ then y\<^sub>\ else y\<^sub>\')}" by(simp add: etp.uni_set_samp Let_def split_def cong: bind_spmf_cong) also have "... = funct_OT_12 x y \ (\ (s1, s2). S1 input\<^sub>1 s1)" by(cases \; simp add: S1_def R1_def Let_def funct_OT_12_def) ultimately show ?thesis by auto qed -text \The adversary used in proof of security for party 2\ +text \ The adversary used in proof of security for party 2 \ definition \ :: "('index, 'range) advP2" where "\ \ \ b\<^sub>\ D2 x = do { \\<^sub>\' \ coin_spmf; x\<^sub>\ \ etp.S \; let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; d \ D2(\, \, \\<^sub>\, \\<^sub>\'); return_spmf(if d then \\<^sub>\' else \ \\<^sub>\')}" lemma lossless_\: assumes "\ view. lossless_spmf (D2 view)" shows "y \ set_spmf I \ lossless_spmf (\ (fst y) \ b\<^sub>\ D2 x)" by(simp add: \_def etp.lossless_S assms) -end - -locale ETP_ot_12 = ETP_base + - fixes HCP_ad :: real -begin - lemma assm_bound_funct_OT_12: assumes "etp.HCP_adv \ \ (if \ then b1 else b0) D \ HCP_ad" shows "\spmf (funct_OT_12 (b0,b1) \ \ (\ (out1,out2). etp.HCP_game \ \ out2 D)) True - 1/2\ \ HCP_ad" (is "?lhs \ HCP_ad") proof- have "?lhs = \spmf (etp.HCP_game \ \ (if \ then b1 else b0) D) True - 1/2\" by(simp add: funct_OT_12_def) thus ?thesis using assms etp.HCP_adv_def by simp qed lemma assm_bound_funct_OT_12_collapse: assumes "\ b\<^sub>\. etp.HCP_adv \ \ b\<^sub>\ D \ HCP_ad" shows "\spmf (funct_OT_12 m1 \ \ (\ (out1,out2). etp.HCP_game \ \ out2 D)) True - 1/2\ \ HCP_ad" using assm_bound_funct_OT_12 surj_pair assms by metis -text \To prove security for party 2 we split the proof on the cases on party 2's input\ +text \ To prove security for party 2 we split the proof on the cases on party 2's input \ lemma R2_S2_False: assumes "((if \ then b0 else b1) = False)" - shows "spmf (R2 (b0,b1) \ \ (D2 :: (bool \ 'a \ bool \ bool) \ bool spmf)) True + shows "spmf (R2 (b0,b1) \ \ (D2 :: (bool \ 'index \ bool \ bool) \ bool spmf)) True = spmf (funct_OT_12 (b0,b1) \ \ (\ (out1,out2). S2 \ out2 \ D2)) True" proof- have "\ \ \ b0" using assms by simp moreover have "\ \ \ \ b1" using assms by simp ultimately show ?thesis - by(auto simp add: R2_def S2_def split_def etp.F_def etp.F\<^sub>i\<^sub>n\<^sub>v_def assms funct_OT_12_def cong: bind_spmf_cong_simp) + by(auto simp add: R2_def S2_def split_def local.etp.F_f_inv assms funct_OT_12_def cong: bind_spmf_cong_simp) qed lemma R2_S2_True: assumes "((if \ then b0 else b1) = True)" and lossless_D: "\ a. lossless_spmf (D2 a)" shows "\(spmf (bind_spmf (R2 (b0,b1) \) D2) True) - spmf (funct_OT_12 (b0,b1) \ \ (\ (out1, out2). S2 \ out2 \ (\ view. D2 view))) True\ = \2*((spmf (etp.HCP_game \ \ (if \ then b1 else b0) D2) True) - 1/2)\" proof- have "(spmf (funct_OT_12 (b0,b1) \ \ (\ (out1, out2). S2 \ out2 \ D2)) True - spmf (bind_spmf (R2 (b0,b1) \) D2) True) = 2 * ((spmf (etp.HCP_game \ \ (if \ then b1 else b0) D2) True) - 1/2)" proof- have "((spmf (etp.HCP_game \ \ (if \ then b1 else b0) D2) True) - 1/2) = 1/2*(spmf (bind_spmf (S2 \ (if \ then b1 else b0)) D2) True - spmf (bind_spmf (R2 (b0,b1) \) D2) True)" including monad_normalisation proof- have \_true_b0_true: "\ \ b0 = True" using assms(1) by simp have \_false_b1_true: "\ \ \ b1" using assms(1) by simp have return_True_False: "spmf (return_spmf (\ d)) True = spmf (return_spmf d) False" for d by(cases d; simp) define HCP_game_true where "HCP_game_true == \ \ b\<^sub>\. do { (\, \) \ I; x\<^sub>\ \ etp.S \; x \ (etp.S \); let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; - let \\<^sub>\' = B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let \\<^sub>\' = B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); d \ D2(\, \, \\<^sub>\, \\<^sub>\'); let b' = (if d then \\<^sub>\' else \ \\<^sub>\'); - let b = B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let b = B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); return_spmf (b = b')}" define HCP_game_false where "HCP_game_false == \ \ b\<^sub>\. do { (\, \) \ I; x\<^sub>\ \ etp.S \; x \ (etp.S \); let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; - let \\<^sub>\' = \ B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let \\<^sub>\' = \ B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); d \ D2(\, \, \\<^sub>\, \\<^sub>\'); let b' = (if d then \\<^sub>\' else \ \\<^sub>\'); - let b = B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let b = B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); return_spmf (b = b')}" define HCP_game_\ where "HCP_game_\ == \ \ b\<^sub>\. do { \\<^sub>\' \ coin_spmf; (\, \) \ I; x \ etp.S \; x' \ etp.S \; d \ D2 (\, \, (B \ x) \ b\<^sub>\, \\<^sub>\'); let b' = (if d then \\<^sub>\' else \ \\<^sub>\'); - return_spmf (B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x') = b')}" + return_spmf (B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x') = b')}" define S2D where "S2D == \ \ b\<^sub>\ . do { (\, \) \ I; x\<^sub>\ \ etp.S \; y\<^sub>\' \ etp.S \; - let x\<^sub>\' = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; + let x\<^sub>\' = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; let \\<^sub>\' = B \ x\<^sub>\'; d :: bool \ D2(\, \, \\<^sub>\, \\<^sub>\'); return_spmf d}" define R2D where "R2D == \ msgs \. do { let (b0,b1) = msgs; (\, \) \ I; x\<^sub>\ \ etp.S \; y\<^sub>\' \ etp.S \; - let y\<^sub>\ = etp.F \ x\<^sub>\; - let x\<^sub>\ = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; - let x\<^sub>\' = etp.F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; + let y\<^sub>\ = F \ x\<^sub>\; + let x\<^sub>\ = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\; + let x\<^sub>\' = F\<^sub>i\<^sub>n\<^sub>v \ \ y\<^sub>\'; let \\<^sub>\ = (B \ x\<^sub>\) \ (if \ then b1 else b0) ; let \\<^sub>\' = (B \ x\<^sub>\') \ (if \ then b0 else b1); b :: bool \ D2(\, \,(\\<^sub>\, \\<^sub>\')); return_spmf b}" define D_true where "D_true == \\ b\<^sub>\. do { (\, \) \ I; x\<^sub>\ \ etp.S \; x \ (etp.S \); let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; - let \\<^sub>\' = B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let \\<^sub>\' = B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); d :: bool \ D2(\, \, \\<^sub>\, \\<^sub>\'); return_spmf d}" define D_false where "D_false == \ \ b\<^sub>\. do { (\, \) \ I; x\<^sub>\ \ etp.S \; x \ etp.S \; let \\<^sub>\ = (B \ x\<^sub>\) \ b\<^sub>\; - let \\<^sub>\' = \ B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x); + let \\<^sub>\' = \ B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x); d :: bool \ D2(\, \, \\<^sub>\, \\<^sub>\'); return_spmf d}" have lossless_D_false: "lossless_spmf (D_false \ (if \ then b1 else b0))" - apply(auto simp add: D_false_def lossless_D lossless_I) + apply(auto simp add: D_false_def lossless_D local.etp.lossless_I) using local.etp.lossless_S by auto have "spmf (etp.HCP_game \ \ (if \ then b1 else b0) D2) True = spmf (HCP_game_\ \ (if \ then b1 else b0)) True" apply(simp add: etp.HCP_game_def HCP_game_\_def \_def split_def etp.F_f_inv) by(rewrite bind_commute_spmf[where q = "coin_spmf"]; rewrite bind_commute_spmf[where q = "coin_spmf"]; rewrite bind_commute_spmf[where q = "coin_spmf"]; auto)+ also have "... = spmf (bind_spmf (map_spmf Not coin_spmf) (\b. if b then HCP_game_true \ (if \ then b1 else b0) else HCP_game_false \ (if \ then b1 else b0))) True" unfolding HCP_game_\_def HCP_game_true_def HCP_game_false_def \_def Let_def apply(simp add: split_def cong: if_cong) supply [[simproc del: monad_normalisation]] apply(subst if_distrib[where f = "bind_spmf _" for f, symmetric]; simp cong: bind_spmf_cong add: if_distribR )+ apply(rewrite in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(fold map_spmf_conv_bind_spmf) apply(rule conjI; rule impI; simp) apply(simp only: spmf_bind) apply(rule Bochner_Integration.integral_cong[OF refl])+ apply clarify subgoal for r r\<^sub>\ \ \ apply(simp only: UNIV_bool spmf_of_set integral_spmf_of_set) apply(simp cong: if_cong split del: if_split) - apply(cases "B r (etp.f\<^sub>i\<^sub>n\<^sub>v r \)") + apply(cases "B r (F\<^sub>i\<^sub>n\<^sub>v r r\<^sub>\ \)") by auto apply(rewrite in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(simp only: spmf_bind) apply(rule Bochner_Integration.integral_cong[OF refl])+ apply clarify subgoal for r r\<^sub>\ \ \ apply(simp only: UNIV_bool spmf_of_set integral_spmf_of_set) apply(simp cong: if_cong split del: if_split) - apply(cases "B r (etp.f\<^sub>i\<^sub>n\<^sub>v r \)") + apply(cases " B r (F\<^sub>i\<^sub>n\<^sub>v r r\<^sub>\ \)") by auto done also have "... = 1/2*(spmf (HCP_game_true \ (if \ then b1 else b0)) True) + 1/2*(spmf (HCP_game_false \ (if \ then b1 else b0)) True)" by(simp add: spmf_bind UNIV_bool spmf_of_set integral_spmf_of_set) also have "... = 1/2*(spmf (D_true \ (if \ then b1 else b0)) True) + 1/2*(spmf (D_false \ (if \ then b1 else b0)) False)" proof- - have "spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x)) \ (\d. return_spmf (\ d)))))) True - = spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x)))))) False" + have "spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x)) \ (\d. return_spmf (\ d)))))) True + = spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x)))))) False" (is "?lhs = ?rhs") proof- - have "?lhs = spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (etp.f\<^sub>i\<^sub>n\<^sub>v \ x)) \ (\d. return_spmf (d)))))) False" + have "?lhs = spmf (I \ (\(\, \). etp.S \ \ (\x\<^sub>\. etp.S \ \ (\x. D2 (\, \, B \ x\<^sub>\ = (\ (if \ then b1 else b0)), \ B \ (F\<^sub>i\<^sub>n\<^sub>v \ \ x)) \ (\d. return_spmf (d)))))) False" by(simp only: split_def return_True_False spmf_bind) then show ?thesis by simp qed then show ?thesis by(simp add: HCP_game_true_def HCP_game_false_def Let_def D_true_def D_false_def if_distrib[where f="(=) _"] cong: if_cong) qed also have "... = 1/2*((spmf (D_true \ (if \ then b1 else b0) ) True) + (1 - spmf (D_false \ (if \ then b1 else b0) ) True))" by(simp add: spmf_False_conv_True lossless_D_false) also have "... = 1/2 + 1/2* (spmf (D_true \ (if \ then b1 else b0)) True) - 1/2*(spmf (D_false \ (if \ then b1 else b0)) True)" by(simp) also have "... = 1/2 + 1/2* (spmf (S2D \ (if \ then b1 else b0) ) True) - 1/2*(spmf (R2D (b0,b1) \ ) True)" - apply(auto simp add: etp.F\<^sub>i\<^sub>n\<^sub>v_def etp.f\<^sub>i\<^sub>n\<^sub>v_def etp.F_def S2D_def R2D_def D_true_def D_false_def assms split_def cong: bind_spmf_cong_simp) + apply(auto simp add: local.etp.F_f_inv S2D_def R2D_def D_true_def D_false_def assms split_def cong: bind_spmf_cong_simp) apply(simp add: \_true_b0_true) by(simp add: \_false_b1_true) ultimately show ?thesis by(simp add: S2D_def R2D_def R2_def S2_def split_def) qed then show ?thesis by(auto simp add: funct_OT_12_def) qed thus ?thesis by simp qed lemma P2_adv_bound: assumes lossless_D: "\ a. lossless_spmf (D2 a)" shows "\(spmf (bind_spmf (R2 (b0,b1) \) D2) True) - spmf (funct_OT_12 (b0,b1) \ \ (\ (out1, out2). S2 \ out2 \ (\ view. D2 view))) True\ \ \2*((spmf (etp.HCP_game \ \ (if \ then b1 else b0) D2) True) - 1/2)\" by(cases "(if \ then b0 else b1)"; auto simp add: R2_S2_False R2_S2_True assms) -sublocale OT_12: sim_det_def R1 S1 R2 S2 funct_OT_12 protocol - unfolding sim_det_def_def by(simp add: lossless_R1 lossless_S1 lossless_funct_OT_12 lossless_R2 lossless_S2) +sublocale OT_12: sim_det_def R1 S1 R2 S2 funct_OT_12 protocol + unfolding sim_det_def_def + by(simp add: lossless_R1 lossless_S1 lossless_R2 lossless_S2 funct_OT_12_def) lemma correct: "OT_12.correctness m1 m2" unfolding OT_12.correctness_def by (metis prod.collapse correctness) -lemma P1_security: "OT_12.inf_theoretic_P1 m1 m2" - unfolding OT_12.inf_theoretic_P1_def using P1_security by simp +lemma P1_security_inf_the: "OT_12.perfect_sec_P1 m1 m2" + unfolding OT_12.perfect_sec_P1_def using P1_security by simp lemma P2_security: assumes "\ a. lossless_spmf (D a)" and "\ b\<^sub>\. etp.HCP_adv \ m2 b\<^sub>\ D \ HCP_ad" shows "OT_12.adv_P2 m1 m2 D \ 2 * HCP_ad" proof- have "spmf (etp.HCP_game \ \ (if \ then b1 else b0) D) True = spmf (funct_OT_12 (b0,b1) \ \ (\ (out1, out2). etp.HCP_game \ \ out2 D)) True" for \ b0 b1 by(simp add: funct_OT_12_def) hence "OT_12.adv_P2 m1 m2 D \ \2*((spmf (funct_OT_12 m1 m2 \ (\ (out1, out2). etp.HCP_game \ m2 out2 D)) True) - 1/2)\" unfolding OT_12.adv_P2_def using P2_adv_bound assms surj_pair prod.collapse by metis moreover have "\2*((spmf (funct_OT_12 m1 m2 \ (\ (out1, out2). etp.HCP_game \ m2 out2 D)) True) - 1/2)\ \ \2*HCP_ad\" proof - have "(\r. \(1::real) / r\ \ 1 / \r\) \ 2 / \1 / (spmf (funct_OT_12 m1 m2 \ (\(x, y). ((\u b. etp.HCP_game \ m2 b D)::unit \ bool \ bool spmf) x y)) True - 1 / 2)\ \ HCP_ad / (1 / 2)" using assm_bound_funct_OT_12_collapse assms by auto then show ?thesis by fastforce qed moreover have "HCP_ad \ 0" using assms(2) local.etp.HCP_adv_def by auto ultimately show ?thesis by argo qed end -text \We also consider the asymptotic case for security proofs\ +text \ We also consider the asymptotic case for security proofs \ locale ETP_sec_para = fixes I :: "nat \ ('index \ 'trap) spmf" and domain :: "'index \ 'range set" and range :: "'index \ 'range set" and f :: "'index \ ('range \ 'range)" + and F :: "'index \ 'range \ 'range" + and F\<^sub>i\<^sub>n\<^sub>v :: "'index \ 'trap \ 'range \ 'range" and B :: "'index \ 'range \ bool" - and etp_advantage :: "nat \ real" - assumes ETP_base: "\ n. ETP_ot_12 (I n) domain range f" + assumes ETP_base: "\ n. ETP_base (I n) domain range F F\<^sub>i\<^sub>n\<^sub>v" begin -sublocale ETP_ot_12 "(I n)" domain range f B "(etp_advantage n)" - using local.ETP_base by simp +sublocale ETP_base "(I n)" domain range + using ETP_base by simp lemma correct_asym: "OT_12.correctness n m1 m2" by(simp add: correct) -lemma P1_sec_asym: "OT_12.inf_theoretic_P1 n m1 m2" - using P1_security by simp +lemma P1_sec_asym: "OT_12.perfect_sec_P1 n m1 m2" + using P1_security_inf_the by simp lemma P2_sec_asym: assumes "\ a. lossless_spmf (D a)" and HCP_adv_neg: "negligible (\ n. etp_advantage n)" and etp_adv_bound: "\ b\<^sub>\ n. etp.HCP_adv n \ m2 b\<^sub>\ D \ etp_advantage n" shows "negligible (\ n. OT_12.adv_P2 n m1 m2 D)" proof- have "negligible (\ n. 2 * etp_advantage n)" using HCP_adv_neg by (simp add: negligible_cmultI) moreover have "\OT_12.adv_P2 n m1 m2 D\ = OT_12.adv_P2 n m1 m2 D" for n unfolding OT_12.adv_P2_def by simp moreover have "OT_12.adv_P2 n m1 m2 D \ 2 * etp_advantage n" for n using assms P2_security by blast ultimately show ?thesis using assms negligible_le HCP_adv_neg P2_security by presburger qed end end \ No newline at end of file diff --git a/thys/Multi_Party_Computation/ETP_RSA_OT.thy b/thys/Multi_Party_Computation/ETP_RSA_OT.thy --- a/thys/Multi_Party_Computation/ETP_RSA_OT.thy +++ b/thys/Multi_Party_Computation/ETP_RSA_OT.thy @@ -1,552 +1,705 @@ -subsubsection \RSA instantiation\ +subsubsection \ RSA instantiation \ text\It is known that the RSA collection forms an ETP. Here we instantitate our proof of security for OT that uses a general ETP for RSA. We use the proof of the general construction of OT. The main proof effort here is in showing the RSA collection meets the requirements of an ETP, mainly this involves showing the RSA mapping is a bijection.\ theory ETP_RSA_OT imports ETP_OT Number_Theory_Aux + Uniform_Sampling begin type_synonym index = "(nat \ nat)" type_synonym trap = nat type_synonym range = nat type_synonym domain = nat type_synonym viewP1 = "((bool \ bool) \ nat \ nat) spmf" type_synonym viewP2 = "(bool \ index \ (bool \ bool)) spmf" type_synonym dist2 = "(bool \ index \ bool \ bool) \ bool spmf" type_synonym advP2 = "index \ bool \ bool \ dist2 \ bool spmf" locale rsa_base = fixes prime_set :: "nat set" \ \the set of primes used\ and B :: "index \ nat \ bool" assumes prime_set_ass: "prime_set \ {x. prime x \ x > 2}" and finite_prime_set: "finite prime_set" and prime_set_gt_2: "card prime_set > 2" begin lemma prime_set_non_empty: "prime_set \ {}" using prime_set_gt_2 by auto definition coprime_set :: "nat \ nat set" where "coprime_set N \ {x. coprime x N \ x > 1 \ x < N}" lemma coprime_set_non_empty: assumes "N > 2" shows "coprime_set N \ {}" by(simp add: coprime_set_def; metis assms(1) Suc_lessE coprime_Suc_right_nat lessI numeral_2_eq_2) definition sample_coprime :: "nat \ nat spmf" where "sample_coprime N = spmf_of_set (coprime_set (N))" +lemma sample_coprime_e_gt_1: + assumes "e \ set_spmf (sample_coprime N)" + shows "e > 1" + using assms by(simp add: sample_coprime_def coprime_set_def) + lemma lossless_sample_coprime: assumes "\ prime N" and "N > 2" shows "lossless_spmf (sample_coprime N)" proof- have "coprime_set N \ {}" by(simp add: coprime_set_non_empty assms) also have "finite (coprime_set N)" by(simp add: coprime_set_def) ultimately show ?thesis by(simp add: sample_coprime_def) qed lemma set_spmf_sample_coprime: shows "set_spmf (sample_coprime N) = {x. coprime x N \ x > 1 \ x < N}" by(simp add: sample_coprime_def coprime_set_def) definition sample_primes :: "nat spmf" where "sample_primes = spmf_of_set prime_set" lemma lossless_sample_primes: shows "lossless_spmf sample_primes" by(simp add: sample_primes_def prime_set_non_empty finite_prime_set) -definition sample_primes_excl :: "nat set \ nat spmf" - where "sample_primes_excl P = spmf_of_set (prime_set - P)" - -lemma lossless_sample_primes_excl: - shows "lossless_spmf (sample_primes_excl {P})" - apply(simp add: sample_primes_excl_def finite_prime_set) - using prime_set_gt_2 subset_singletonD by fastforce - lemma set_spmf_sample_primes: shows "set_spmf sample_primes \ {x. prime x \ x > 2}" by(auto simp add: sample_primes_def prime_set_ass finite_prime_set) lemma mem_samp_primes_gt_2: shows "x \ set_spmf sample_primes \ x > 2" apply (simp add: finite_prime_set sample_primes_def) using prime_set_ass by blast lemma mem_samp_primes_prime: shows "x \ set_spmf sample_primes \ prime x" apply (simp add: finite_prime_set sample_primes_def prime_set_ass) using prime_set_ass by blast +definition sample_primes_excl :: "nat set \ nat spmf" + where "sample_primes_excl P = spmf_of_set (prime_set - P)" + +lemma lossless_sample_primes_excl: + shows "lossless_spmf (sample_primes_excl {P})" + apply(simp add: sample_primes_excl_def finite_prime_set) + using prime_set_gt_2 subset_singletonD by fastforce + +definition sample_set_excl :: "nat set \ nat set \ nat spmf" + where "sample_set_excl Q P = spmf_of_set (Q - P)" + +lemma set_spmf_sample_set_excl [simp]: + assumes "finite (Q - P)" + shows "set_spmf (sample_set_excl Q P) = (Q - P)" + unfolding sample_set_excl_def + by (metis set_spmf_of_set assms)+ + +lemma lossless_sample_set_excl: + assumes "finite Q" + and "card Q > 2" + shows "lossless_spmf (sample_set_excl Q {P})" + unfolding sample_set_excl_def + using assms subset_singletonD by fastforce + lemma mem_samp_primes_excl_gt_2: - shows "x \ set_spmf (sample_primes_excl {y}) \ x > 2" - apply(simp add: finite_prime_set sample_primes_excl_def prime_set_ass ) + shows "x \ set_spmf (sample_set_excl prime_set {y}) \ x > 2" + apply(simp add: finite_prime_set sample_set_excl_def prime_set_ass ) using prime_set_ass by blast lemma mem_samp_primes_excl_prime : - shows "x \ set_spmf (sample_primes_excl {y}) \ prime x" - apply (simp add: finite_prime_set sample_primes_excl_def) + shows "x \ set_spmf (sample_set_excl prime_set {y}) \ prime x" + apply (simp add: finite_prime_set sample_set_excl_def) using prime_set_ass by blast lemma sample_coprime_lem: assumes "x \ set_spmf sample_primes" - and " y \ set_spmf (sample_primes_excl {x}) " + and " y \ set_spmf (sample_set_excl prime_set {x}) " shows "lossless_spmf (sample_coprime ((x - Suc 0) * (y - Suc 0)))" proof- have gt_2: "x > 2" "y > 2" using mem_samp_primes_gt_2 assms mem_samp_primes_excl_gt_2 by auto have "\ prime ((x-1)*(y-1))" proof- have "prime x" "prime y" using mem_samp_primes_prime mem_samp_primes_excl_prime assms by auto then show ?thesis using prod_not_prime gt_2 by simp qed also have "((x-1)*(y-1)) > 2" by (metis (no_types, lifting) gt_2 One_nat_def Suc_diff_1 assms(1) assms(2) calculation divisors_zero less_2_cases nat_1_eq_mult_iff nat_neq_iff not_numeral_less_one numeral_2_eq_2 prime_gt_0_nat rsa_base.mem_samp_primes_excl_prime rsa_base.mem_samp_primes_prime rsa_base_axioms two_is_prime_nat) ultimately show ?thesis using lossless_sample_coprime by simp qed definition I :: "(index \ trap) spmf" where "I = do { P \ sample_primes; - Q \ sample_primes_excl {P}; + Q \ sample_set_excl prime_set {P}; let N = P*Q; let N' = (P-1)*(Q-1); e \ sample_coprime N'; let d = nat ((fst (bezw e N')) mod N'); return_spmf ((N, e), d)}" lemma lossless_I: "lossless_spmf I" - by(auto simp add: I_def lossless_sample_primes lossless_sample_primes_excl Let_def sample_coprime_lem) + by(auto simp add: I_def lossless_sample_primes lossless_sample_set_excl finite_prime_set prime_set_gt_2 Let_def sample_coprime_lem) -lemma set_spmf_I: +lemma set_spmf_I_N: assumes "((N,e),d) \ set_spmf I" - obtains P Q where "N = P * Q" and "P \ Q" and "prime P" and "prime Q" and "coprime e ((P - 1)*(Q - 1))" + obtains P Q where "N = P * Q" + and "P \ Q" + and "prime P" + and "prime Q" + and "coprime e ((P - 1)*(Q - 1))" + and "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" using assms apply(auto simp add: I_def Let_def) - using finite_prime_set mem_samp_primes_prime rsa_base.sample_primes_excl_def rsa_base_axioms sample_primes_def + using finite_prime_set mem_samp_primes_prime sample_set_excl_def rsa_base_axioms sample_primes_def by (simp add: set_spmf_sample_coprime) +lemma set_spmf_I_e_d: + assumes "((N,e),d) \ set_spmf I" + shows "e > 1" and "d > 1" + using assms sample_coprime_e_gt_1 + apply(auto simp add: I_def Let_def) + by (smt Euclidean_Division.pos_mod_sign Num.of_nat_simps(5) Suc_diff_1 bezw_inverse cong_def coprime_imp_gcd_eq_1 gr0I less_1_mult less_numeral_extra(2) mem_Collect_eq mod_by_0 mod_less more_arith_simps(6) nat_0 nat_0_less_mult_iff nat_int nat_neq_iff numerals(2) of_nat_0_le_iff of_nat_1 rsa_base.mem_samp_primes_gt_2 rsa_base_axioms set_spmf_sample_coprime zero_less_diff) + definition domain :: "index \ nat set" where "domain index \ {..< fst index}" definition range :: "index \ nat set" where "range index \ {..< fst index}" lemma finite_range: "finite (range index)" by(simp add: range_def) lemma dom_eq_ran: "domain index = range index" by(simp add: range_def domain_def) -definition f :: "index \ (nat \ nat)" - where "f index x = x ^ (snd index) mod (fst index)" +definition F :: "index \ (nat \ nat)" + where "F index x = x ^ (snd index) mod (fst index)" -text \We must prove the RSA function is a bijection\ +definition F\<^sub>i\<^sub>n\<^sub>v :: "index \ trap \ nat \ nat" + where "F\<^sub>i\<^sub>n\<^sub>v \ \ y = y ^ \ mod (fst \)" + +text \ We must prove the RSA function is a bijection \ lemma rsa_bijection: assumes coprime: "coprime e ((P-1)*(Q-1))" and prime_P: "prime (P::nat)" and prime_Q: "prime Q" and P_neq_Q: "P \ Q" and x_lt_pq: "x < P * Q" and y_lt_pd: "y < P * Q" and rsa_map_eq: "x ^ e mod (P * Q) = y ^ e mod (P * Q)" shows "x = y" proof- have flt_xP: "[x ^ P = x] (mod P)" using fermat_little prime_P by blast have flt_yP: "[y ^ P = y] (mod P)" using fermat_little prime_P by blast have flt_xQ: "[x ^ Q = x] (mod Q)" using fermat_little prime_Q by blast have flt_yQ: "[y ^ Q = y] (mod Q)" using fermat_little prime_Q by blast show ?thesis proof(cases "y \ x") case True hence ye_gt_xe: "y^e \ x^e" by (simp add: power_mono) have x_y_exp_e: "[x^e = y^e] (mod P)" using cong_modulus_mult_nat cong_altdef_nat True ye_gt_xe cong_sym cong_def assms by blast obtain d where d: "[e*d = 1] (mod (P-1)) \ d \ 0" using ex_inverse assms by blast then obtain k where k: "e*d = 1 + k*(P-1)" using ex_k_mod assms by blast hence xk_yk: "[x^(1 + k*(P-1)) = y^(1 + k*(P-1))] (mod P)" by(metis k power_mult x_y_exp_e cong_pow) have xk_x: "[x^(1 + k*(P-1)) = x] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[x ^ (1 + k * (P - 1)) = x] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k * (P - 1)) = x] (mod P)" using asm by simp hence "[x ^ (k * (P - 1)) * x ^ P = x] (mod P)" using flt_xP by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k * (P - 1) + P) = x] (mod P)" by (simp add: power_add) hence "[x ^ (1 + (k + 1) * (P - 1)) = x] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y: "[y^(1 + k*(P-1)) = y] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[y ^ (1 + k * (P - 1)) = y] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k * (P - 1)) = y] (mod P)" using asm by simp hence "[y ^ (k * (P - 1)) * y ^ P = y] (mod P)" using flt_yP by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k * (P - 1) + P) = y] (mod P)" by (simp add: power_add) hence "[y ^ (1 + (k + 1) * (P - 1)) = y] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have "[x^e = y^e] (mod Q)" by (metis rsa_map_eq cong_modulus_mult_nat cong_def mult.commute) obtain d' where d': "[e*d' = 1] (mod (Q-1)) \ d' \ 0" by (metis mult.commute ex_inverse prime_P prime_Q P_neq_Q coprime) then obtain k' where k': "e*d' = 1 + k'*(Q-1)" by(metis ex_k_mod mult.commute prime_P prime_Q P_neq_Q coprime) hence xk_yk': "[x^(1 + k'*(Q-1)) = y^(1 + k'*(Q-1))] (mod Q)" by(metis k' power_mult \[x ^ e = y ^ e] (mod Q)\ cong_pow) have xk_x': "[x^(1 + k'*(Q-1)) = x] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[x ^ (1 + k' * (Q - 1)) = x] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k' * (Q - 1)) = x] (mod Q)" using asm by simp hence "[x ^ (k' * (Q - 1)) * x ^ Q = x] (mod Q)" using flt_xQ by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k' * (Q - 1) + Q) = x] (mod Q)" by (simp add: power_add) hence "[x ^ (1 + (k' + 1) * (Q - 1)) = x] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y': "[y^(1 + k'*(Q-1)) = y] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[y ^ (1 + k' * (Q - 1)) = y] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k' * (Q - 1)) = y] (mod Q)" using asm by simp hence "[y ^ (k' * (Q - 1)) * y ^ Q = y] (mod Q)" using flt_yQ by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k' * (Q - 1) + Q) = y] (mod Q)" by (simp add: power_add) hence "[y ^ (1 + (k' + 1) * (Q - 1)) = y] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have P_dvd_xy: "P dvd (y - x)" proof- have "[x = y] (mod P)" using xk_x yk_y xk_yk by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym True by blast qed have Q_dvd_xy: "Q dvd (y - x)" proof- have "[x = y] (mod Q)" using xk_x' yk_y' xk_yk' by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym True by blast qed show ?thesis proof- have "P*Q dvd (y - x)" using P_dvd_xy Q_dvd_xy by (simp add: assms divides_mult primes_coprime) then have "[x = y] (mod P*Q)" by (simp add: cong_altdef_nat cong_sym True) hence "x mod P*Q = y mod P*Q" using cong_def xk_x xk_yk yk_y by metis then show ?thesis using \[x = y] (mod P * Q)\ cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast qed next case False hence ye_gt_xe: "x^e \ y^e" by (simp add: power_mono) have pow_eq: "[x^e = y^e] (mod P*Q)" by (simp add: cong_def assms) then have PQ_dvd_ye_xe: "(P*Q) dvd (x^e - y^e)" using cong_altdef_nat False ye_gt_xe cong_sym by blast then have "[x^e = y^e] (mod P)" using cong_modulus_mult_nat pow_eq by blast obtain d where d: "[e*d = 1] (mod (P-1)) \ d \ 0" using ex_inverse assms by blast then obtain k where k: "e*d = 1 + k*(P-1)" using ex_k_mod assms by blast have xk_yk: "[x^(1 + k*(P-1)) = y^(1 + k*(P-1))] (mod P)" proof- have "[(x^e)^d = (y^e)^d] (mod P)" using \[x ^ e = y ^ e] (mod P)\ cong_pow by blast then have "[x^(e*d) = y^(e*d)] (mod P)" by (simp add: power_mult) thus ?thesis using k by simp qed have xk_x: "[x^(1 + k*(P-1)) = x] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[x ^ (1 + k * (P - 1)) = x] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k * (P - 1)) = x] (mod P)" using asm by simp hence "[x ^ (k * (P - 1)) * x ^ P = x] (mod P)" using flt_xP by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k * (P - 1) + P) = x] (mod P)" by (simp add: power_add) hence "[x ^ (1 + (k + 1) * (P - 1)) = x] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y: "[y^(1 + k*(P-1)) = y] (mod P)" proof(induct k) case 0 then show ?case by simp next case (Suc k) assume asm: "[y ^ (1 + k * (P - 1)) = y] (mod P)" then show ?case proof- have exp_rewrite: "(k * (P - 1) + P) = (1 + (k + 1) * (P - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_P prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k * (P - 1)) = y] (mod P)" using asm by simp hence "[y ^ (k * (P - 1)) * y ^ P = y] (mod P)" using flt_yP by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k * (P - 1) + P) = y] (mod P)" by (simp add: power_add) hence "[y ^ (1 + (k + 1) * (P - 1)) = y] (mod P)" using exp_rewrite by argo thus ?thesis by simp qed qed have P_dvd_xy: "P dvd (x - y)" proof- have "[x = y] (mod P)" using xk_x yk_y xk_yk by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym False by simp qed have "[x^e = y^e] (mod Q)" using cong_modulus_mult_nat pow_eq PQ_dvd_ye_xe cong_dvd_modulus_nat dvd_triv_right by blast obtain d' where d': "[e*d' = 1] (mod (Q-1)) \ d' \ 0" by (metis mult.commute ex_inverse prime_P prime_Q coprime P_neq_Q) then obtain k' where k': "e*d' = 1 + k'*(Q-1)" by(metis ex_k_mod mult.commute prime_P prime_Q coprime P_neq_Q) have xk_yk': "[x^(1 + k'*(Q-1)) = y^(1 + k'*(Q-1))] (mod Q)" proof- have "[(x^e)^d' = (y^e)^d'] (mod Q)" using \[x ^ e = y ^ e] (mod Q)\ cong_pow by blast then have "[x^(e*d') = y^(e*d')] (mod Q)" by (simp add: power_mult) thus ?thesis using k' by simp qed have xk_x': "[x^(1 + k'*(Q-1)) = x] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[x ^ (1 + k' * (Q - 1)) = x] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[x * x ^ (k' * (Q - 1)) = x] (mod Q)" using asm by simp hence "[x ^ (k' * (Q - 1)) * x ^ Q = x] (mod Q)" using flt_xQ by (metis cong_scalar_right cong_trans mult.commute) hence "[x ^ (k' * (Q - 1) + Q) = x] (mod Q)" by (simp add: power_add) hence "[x ^ (1 + (k' + 1) * (Q - 1)) = x] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have yk_y': "[y^(1 + k'*(Q-1)) = y] (mod Q)" proof(induct k') case 0 then show ?case by simp next case (Suc k') assume asm: "[y ^ (1 + k' * (Q - 1)) = y] (mod Q)" then show ?case proof- have exp_rewrite: "(k' * (Q - 1) + Q) = (1 + (k' + 1) * (Q - 1))" by (smt add.assoc add.commute le_add_diff_inverse nat_le_linear not_add_less1 prime_Q prime_gt_1_nat semiring_normalization_rules(3)) have "[y * y ^ (k' * (Q - 1)) = y] (mod Q)" using asm by simp hence "[y ^ (k' * (Q - 1)) * y ^ Q = y] (mod Q)" using flt_yQ by (metis cong_scalar_right cong_trans mult.commute) hence "[y ^ (k' * (Q - 1) + Q) = y] (mod Q)" by (simp add: power_add) hence "[y ^ (1 + (k' + 1) * (Q - 1)) = y] (mod Q)" using exp_rewrite by argo thus ?thesis by simp qed qed have Q_dvd_xy: "Q dvd (x - y)" proof- have "[x = y] (mod Q)" using xk_x' yk_y' xk_yk' by (simp add: cong_def) thus ?thesis using cong_altdef_nat cong_sym False by simp qed show ?thesis proof- have "P*Q dvd (x - y)" using P_dvd_xy Q_dvd_xy by (simp add: assms divides_mult primes_coprime) hence 1: "[x = y] (mod P*Q)" using False cong_altdef_nat linear by blast hence "x mod P*Q = y mod P*Q" using cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast thus ?thesis using 1 cong_less_modulus_unique_nat x_lt_pq y_lt_pd by blast qed qed qed lemma rsa_bij_betw: assumes "coprime e ((P - 1)*(Q - 1))" and "prime P" and "prime Q" and "P \ Q" - shows "bij_betw (f ((P * Q), e)) (range ((P * Q), e)) (range ((P * Q), e))" + shows "bij_betw (F ((P * Q), e)) (range ((P * Q), e)) (range ((P * Q), e))" proof- have PQ_not_0: "prime P \ prime Q \ P * Q \ 0" using assms by auto have "inj_on (\x. x ^ snd (P * Q, e) mod fst (P * Q, e)) {..x. x ^ snd (P * Q, e) mod fst (P * Q, e)) ` {.. set_spmf I" - shows " bij_betw (f ((N), e)) (range ((N), e)) (range ((N), e))" + shows "bij_betw (F ((N), e)) (range ((N), e)) (range ((N), e))" proof- - obtain P Q where "N = P * Q" and "bij_betw (f ((P*Q), e)) (range ((P*Q), e)) (range ((P*Q), e))" + obtain P Q where "N = P * Q" and "bij_betw (F ((P*Q), e)) (range ((P*Q), e)) (range ((P*Q), e))" proof- obtain P Q where "prime P" and "prime Q" and "N = P * Q" and "P \ Q" and "coprime e ((P - 1)*(Q - 1))" - using set_spmf_I assms by blast + using set_spmf_I_N assms by blast then show ?thesis using rsa_bij_betw that by blast qed thus ?thesis by blast qed -sublocale etp: etp I domain range f B - by (auto simp add: etp_def dom_eq_ran finite_range bij_betw1 lossless_I; metis Iio_eq_empty_iff bot_nat_def fst_conv mult_eq_0_iff not_prime_0 range_def set_spmf_I) +lemma + assumes "P dvd x" +shows "[x = 0] (mod P)" + using assms cong_def by force -sublocale etp: ETP_base I domain range f B - apply(auto simp add: ETP_base_def dom_eq_ran finite_range local.etp.non_empty_range lossless_I) - using dom_eq_ran local.etp.non_empty_range bij_betw1 by auto +lemma rsa_inv: + assumes d: "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" + and coprime: "coprime e ((P-1)*(Q-1))" + and prime_P: "prime (P::nat)" + and prime_Q: "prime Q" + and P_neq_Q: "P \ Q" + and e_gt_1: "e > 1" + and d_gt_1: "d > 1" + shows "((((x) ^ e) mod (P*Q)) ^ d) mod (P*Q) = x mod (P*Q)" +proof(cases "x = 0 \ x = 1") + case True + then show ?thesis + by (metis assms(6) assms(7) le_simps(1) nat_power_eq_Suc_0_iff neq0_conv not_one_le_zero numeral_nat(7) power_eq_0_iff power_mod) +next + case False + hence x_gt_1: "x > 1" by simp + define z where "z = (x ^ e) ^ d - x" + hence z_gt_0: "z > 0" + proof- + have "(x ^ e) ^ d - x = x ^ (e * d) - x" + by (simp add: power_mult) + also have "... > 0" + by (metis x_gt_1 e_gt_1 d_gt_1 le_neq_implies_less less_one linorder_not_less n_less_m_mult_n not_less_zero numeral_nat(7) power_increasing_iff power_one_right zero_less_diff) + ultimately show ?thesis using z_def by argo + qed + hence "[z = 0] (mod P)" + proof(cases "[x = 0] (mod P)") + case True + then show ?thesis + proof - + have "0 \ d * e" + by (metis (no_types) assms assms mult_is_0 not_one_less_zero) + then show ?thesis + by (metis (no_types) Groups.add_ac(2) True add_diff_inverse_nat cong_def cong_dvd_iff cong_mult_self_right dvd_0_right dvd_def dvd_trans mod_add_self2 more_arith_simps(5) nat_diff_split power_eq_if power_mult semiring_normalization_rules(7) z_def) + qed + next + case False + have "[e * d = 1] (mod ((P - 1) * (Q - 1)))" + by (metis d bezw_inverse coprime coprime_imp_gcd_eq_1 nat_int) + hence "[e * d = 1] (mod (P - 1))" + using assms cong_modulus_mult_nat by blast + then obtain k where k: "e*d = 1 + k*(P-1)" + using ex_k_mod assms by force + hence "x ^ (e * d) = x * ((x ^ (P - 1)) ^ k)" + by (metis power_add power_one_right mult.commute power_mult) + hence "[x ^ (e * d) = x * ((x ^ (P - 1)) ^ k)] (mod P)" + using cong_def by simp + moreover have "[x ^ (P - 1) = 1] (mod P)" + using prime_P fermat_theorem False + by (simp add: cong_0_iff) + moreover have "[x ^ (e * d) = x * ((1) ^ k)] (mod P)" + by (metis \x ^ (e * d) = x * (x ^ (P - 1)) ^ k\ calculation(2) cong_pow cong_scalar_left) + hence "[x ^ (e * d) = x] (mod P)" by simp + thus ?thesis using z_def z_gt_0 + by (simp add: cong_diff_iff_cong_0_nat power_mult) + qed + moreover have "[z = 0] (mod Q)" + proof(cases "[x = 0] (mod Q)") + case True + then show ?thesis + proof - + have "0 \ d * e" + by (metis (no_types) assms mult_is_0 not_one_less_zero) + then show ?thesis + by (metis (no_types) Groups.add_ac(2) True add_diff_inverse_nat cong_def cong_dvd_iff cong_mult_self_right dvd_0_right dvd_def dvd_trans mod_add_self2 more_arith_simps(5) nat_diff_split power_eq_if power_mult semiring_normalization_rules(7) z_def) + qed + next + case False + have "[e * d = 1] (mod ((P - 1) * (Q - 1)))" + by (metis d bezw_inverse coprime coprime_imp_gcd_eq_1 nat_int) + hence "[e * d = 1] (mod (Q - 1))" + using assms cong_modulus_mult_nat mult.commute by metis + then obtain k where k: "e*d = 1 + k*(Q-1)" + using ex_k_mod assms by force + hence "x ^ (e * d) = x * ((x ^ (Q - 1)) ^ k)" + by (metis power_add power_one_right mult.commute power_mult) + hence "[x ^ (e * d) = x * ((x ^ (Q - 1)) ^ k)] (mod P)" + using cong_def by simp + moreover have "[x ^ (Q - 1) = 1] (mod Q)" + using prime_Q fermat_theorem False + by (simp add: cong_0_iff) + moreover have "[x ^ (e * d) = x * ((1) ^ k)] (mod Q)" + by (metis \x ^ (e * d) = x * (x ^ (Q - 1)) ^ k\ calculation(2) cong_pow cong_scalar_left) + hence "[x ^ (e * d) = x] (mod Q)" by simp + thus ?thesis using z_def z_gt_0 + by (simp add: cong_diff_iff_cong_0_nat power_mult) + qed + ultimately have "Q dvd (x ^ e) ^ d - x" + "P dvd (x ^ e) ^ d - x" + using z_def assms cong_0_iff by blast + + hence "P * Q dvd ((x ^ e) ^ d - x)" + using assms divides_mult primes_coprime_nat by blast + hence "[(x ^ e) ^ d = x] (mod (P * Q))" + using z_gt_0 cong_altdef_nat z_def by auto + thus ?thesis + by (simp add: cong_def power_mod) +qed -end + +lemma rsa_inv_set_spmf_I: + assumes "((N, e), d) \ set_spmf I" + shows "((((x::nat) ^ e) mod N) ^ d) mod N = x mod N" +proof- + obtain P Q where "N = P * Q" and "d = nat (fst (bezw e ((P-1)*(Q-1))) mod int ((P-1)*(Q-1)))" + and "prime P" + and "prime Q" + and "coprime e ((P - 1)*(Q - 1))" + and "P \ Q" + using assms set_spmf_I_N + by blast + moreover have "e > 1" and "d > 1" using set_spmf_I_e_d assms by auto + ultimately show ?thesis using rsa_inv by blast +qed + +sublocale etp_rsa: etp I domain range F F\<^sub>i\<^sub>n\<^sub>v + unfolding etp_def apply(auto simp add: etp_def dom_eq_ran finite_range bij_betw1 lossless_I) + apply (metis fst_conv lessThan_iff mem_simps(2) nat_0_less_mult_iff prime_gt_0_nat range_def set_spmf_I_N) + apply(auto simp add: F_def F\<^sub>i\<^sub>n\<^sub>v_def) using rsa_inv_set_spmf_I + by (simp add: range_def) + +sublocale etp: ETP_base I domain range B F F\<^sub>i\<^sub>n\<^sub>v + unfolding ETP_base_def + by (simp add: etp_rsa.etp_axioms) + text\After proving the RSA collection is an ETP the proofs of security come easily from the general proofs.\ -locale rsa_proof = rsa_base + - fixes HCP_ad :: real -begin - -sublocale etp_proof: ETP_ot_12 I domain range f - by (auto simp add: etp.ETP_base_axioms ETP_ot_12_def) +lemma correctness_rsa: "etp.OT_12.correctness m1 m2" + by (rule local.etp.correct) -lemma "sim_det_def.correctness funct_OT_12 etp.protocol m1 m2" - by (metis etp.correctness etp_proof.OT_12.correctness_def surj_pair) +lemma P1_security_rsa: "etp.OT_12.perfect_sec_P1 m1 m2" + by(rule local.etp.P1_security_inf_the) -lemma P1_sec: "sim_det_def.inf_theoretic_P1 etp.R1 etp.S1 funct_OT_12 m1 m2" - using local.etp_proof.P1_security by auto - -lemma P2_sec: +lemma P2_security_rsa: assumes "\ a. lossless_spmf (D a)" - and "\b\<^sub>\. local.etp.HCP_adv etp.\ m2 b\<^sub>\ D \ HCP_ad" - shows "sim_det_def.adv_P2 etp.R2 etp.S2 funct_OT_12 m1 m2 D \ 2 * HCP_ad" - using assms etp_proof.P2_security by blast + and "\b\<^sub>\. local.etp_rsa.HCP_adv etp.\ m2 b\<^sub>\ D \ HCP_ad" + shows "etp.OT_12.adv_P2 m1 m2 D \ 2 * HCP_ad" + by(simp add: local.etp.P2_security assms) end locale rsa_asym = fixes prime_set :: "nat \ nat set" and B :: "index \ nat \ bool" - and hcp_advantage :: "nat \ real" - assumes rsa_proof_assm: "\ n. rsa_proof (prime_set n)" + assumes rsa_proof_assm: "\ n. rsa_base (prime_set n)" begin -sublocale rsa_proof "(prime_set n)" B "(hcp_advantage n)" +sublocale rsa_base "(prime_set n)" B using local.rsa_proof_assm by simp -lemma P1_sec_asym: "sim_det_def.inf_theoretic_P1 (etp.R1 n) (etp.S1 n) funct_OT_12 m1 m2" - using P1_sec by simp +lemma correctness_rsa_asymp: + shows "etp.OT_12.correctness n m1 m2" + by(rule correctness_rsa) + +lemma P1_sec_asymp: "etp.OT_12.perfect_sec_P1 n m1 m2" + by(rule local.P1_security_rsa) lemma P2_sec_asym: assumes "\ a. lossless_spmf (D a)" and HCP_adv_neg: "negligible (\ n. hcp_advantage n)" - and hcp_adv_bound: "\b\<^sub>\ n. local.etp.HCP_adv n etp.\ m2 b\<^sub>\ D \ hcp_advantage n" - shows "negligible (\ n. sim_det_def.adv_P2 (etp.R2 n) (etp.S2 n) funct_OT_12 m1 m2 D)" + and hcp_adv_bound: "\b\<^sub>\ n. local.etp_rsa.HCP_adv n etp.\ m2 b\<^sub>\ D \ hcp_advantage n" + shows "negligible (\ n. etp.OT_12.adv_P2 n m1 m2 D)" proof- have "negligible (\ n. 2 * hcp_advantage n)" using HCP_adv_neg by (simp add: negligible_cmultI) - moreover have "\sim_det_def.adv_P2 (etp.R2 n) (etp.S2 n) funct_OT_12 m1 m2 D\ - = sim_det_def.adv_P2 (etp.R2 n) (etp.S2 n) funct_OT_12 m1 m2 D" - for n unfolding sim_det_def.adv_P2_def etp_proof.OT_12.adv_P2_def by linarith - moreover have "sim_det_def.adv_P2 (etp.R2 n) (etp.S2 n) funct_OT_12 m1 m2 D \ 2 * hcp_advantage n" for n - using P2_sec assms by blast + moreover have "\etp.OT_12.adv_P2 n m1 m2 D\ = etp.OT_12.adv_P2 n m1 m2 D" + for n unfolding sim_det_def.adv_P2_def local.etp.OT_12.adv_P2_def by linarith + moreover have "etp.OT_12.adv_P2 n m1 m2 D \ 2 * hcp_advantage n" for n + using P2_security_rsa assms by blast ultimately show ?thesis using assms negligible_le by presburger qed end end \ No newline at end of file diff --git a/thys/Multi_Party_Computation/GMW.thy b/thys/Multi_Party_Computation/GMW.thy --- a/thys/Multi_Party_Computation/GMW.thy +++ b/thys/Multi_Party_Computation/GMW.thy @@ -1,488 +1,488 @@ subsection \1-out-of-4 OT to GMW\ text\We prove security for the gates of the GMW protocol in the semi honest model. We assume security on 1-out-of-4 OT.\ theory GMW imports OT14 begin type_synonym share_1 = bool type_synonym share_2 = bool type_synonym shares_1 = "bool list" type_synonym shares_2 = "bool list" type_synonym msgs_14_OT = "(bool \ bool \ bool \ bool)" type_synonym choice_14_OT = "(bool \ bool)" type_synonym share_wire = "(share_1 \ share_2)" locale gmw_base = fixes S1_14_OT :: "msgs_14_OT \ unit \ 'v_14_OT1 spmf" \ \simulated view for party 1 of OT14\ and R1_14_OT :: "msgs_14_OT \ choice_14_OT \ 'v_14_OT1 spmf" \ \real view for party 1 of OT14\ and S2_14_OT :: "choice_14_OT \ bool \ 'v_14_OT2 spmf" and R2_14_OT :: "msgs_14_OT \ choice_14_OT \ 'v_14_OT2 spmf" and protocol_14_OT :: "msgs_14_OT \ choice_14_OT \ (unit \ bool) spmf" and adv_14_OT :: real assumes P1_OT_14_adv_bound: "sim_det_def.adv_P1 R1_14_OT S1_14_OT funct_14_OT M C D \ adv_14_OT" \ \bound the advantage of party 1 in the 1-out-of-4 OT\ - and P2_OT_12_inf_theoretic: "sim_det_def.inf_theoretic_P2 R2_14_OT S2_14_OT funct_14_OT M C" \ \information theoretic security for party 2 in the 1-out-of-4 OT\ + and P2_OT_12_inf_theoretic: "sim_det_def.perfect_sec_P2 R2_14_OT S2_14_OT funct_14_OT M C" \ \information theoretic security for party 2 in the 1-out-of-4 OT\ and correct_14: "funct_OT_14 msgs C = protocol_14_OT msgs C" \ \correctness of the 1-out-of-4 OT\ and lossless_R1_14_OT: "lossless_spmf (R1_14_OT (m1,m2,m3,m4) (c0,c1))" and lossless_R2_14_OT: "lossless_spmf (R2_14_OT (m1,m2,m3,m4) (c0,c1))" and lossless_S1_14_OT: "lossless_spmf (S1_14_OT (m1,m2,m3,m4) ())" and lossless_S2_14_OT: "lossless_spmf (S2_14_OT (c0,c1) b)" and lossless_protocol_14_OT: "lossless_spmf (protocol_14_OT S C)" and lossless_funct_14_OT: "lossless_spmf (funct_14_OT M C)" begin lemma funct_14: "funct_OT_14 (m00,m01,m10,m11) (c0,c1) = return_spmf ((),if c0 then (if c1 then m11 else m10) else (if c1 then m01 else m00))" by(simp add: funct_OT_14_def) sublocale OT_14_sim: sim_det_def R1_14_OT S1_14_OT R2_14_OT S2_14_OT funct_14_OT protocol_14_OT unfolding sim_det_def_def by(simp add: lossless_R1_14_OT lossless_S1_14_OT lossless_funct_14_OT lossless_R2_14_OT lossless_S2_14_OT) lemma inf_th_14_OT_P4: "R2_14_OT msgs C = (funct_OT_14 msgs C \ (\ (s1, s2). S2_14_OT C s2))" - using P2_OT_12_inf_theoretic sim_det_def.inf_theoretic_P2_def OT_14_sim.inf_theoretic_P2_def by auto + using P2_OT_12_inf_theoretic sim_det_def.perfect_sec_P2_def OT_14_sim.perfect_sec_P2_def by auto lemma ass_adv_14_OT: "\spmf (bind_spmf (S1_14_OT msgs ()) (\ view. (D view))) True - spmf (bind_spmf (R1_14_OT msgs (c0,c1)) (\ view. (D view))) True \ \ adv_14_OT" (is "?lhs \ adv_14_OT") proof- have "funct_OT_14 (m0,m1,m2,m3) (c0, c1) \ (\(o1, o2). S1_14_OT (m0,m1,m2,m3) () \ D) = S1_14_OT (m0,m1,m2,m3) () \ D" for m0 m1 m2 m3 by(simp add: funct_14) hence funct: "funct_OT_14 msgs (c0, c1) \ (\(o1, o2). S1_14_OT msgs () \ D) = S1_14_OT msgs () \ D" by (metis prod_cases4) have "?lhs = \spmf (bind_spmf (R1_14_OT msgs (c0,c1)) (\ view. (D view))) True - spmf (bind_spmf (S1_14_OT msgs ()) (\ view. (D view))) True\" by linarith hence "... = \(spmf (R1_14_OT msgs (c0,c1) \ (\ view. D view)) True) - spmf (funct_OT_14 msgs (c0,c1) \ (\ (o1, o2). S1_14_OT msgs o1 \ (\ view. D view))) True\" by(simp add: funct) thus ?thesis using P1_OT_14_adv_bound sim_det_def.adv_P1_def by (simp add: OT_14_sim.adv_P1_def abs_minus_commute) qed text \The sharing scheme\ definition share :: "bool \ share_wire spmf" where "share x = do { a\<^sub>1 \ coin_spmf; let b\<^sub>1 = x \ a\<^sub>1; return_spmf (a\<^sub>1, b\<^sub>1)}" lemma lossless_share [simp]: "lossless_spmf (share x)" by(simp add: share_def) definition reconstruct :: "(share_1 \ share_2) \ bool spmf" where "reconstruct shares = do { let (a,b) = shares; return_spmf (a \ b)}" lemma lossless_reconstruct [simp]: "lossless_spmf (reconstruct s)" by(simp add: reconstruct_def split_def) lemma reconstruct_share : "(bind_spmf (share x) reconstruct) = (return_spmf x)" proof- have "y = (y = x) = x" for y by auto thus ?thesis by(auto simp add: share_def reconstruct_def bind_spmf_const eq_commute) qed lemma "(reconstruct (s1,s2) \ (\ rec. share rec \ (\ shares. reconstruct shares))) = return_spmf (s1 \ s2)" apply(simp add: reconstruct_share reconstruct_def share_def) apply(cases s1; cases s2) by(auto simp add: bind_spmf_const) definition xor_evaluate :: "bool \ bool \ bool spmf" where "xor_evaluate A B = return_spmf (A \ B)" definition xor_funct :: "share_wire \ share_wire \ (bool \ bool) spmf" where "xor_funct A B = do { let (a1, b1) = A; let (a2,b2) = B; return_spmf (a1 \ a2, b1 \ b2)}" lemma lossless_xor_funct: "lossless_spmf (xor_funct A B)" by(simp add: xor_funct_def split_def) definition xor_protocol :: "share_wire \ share_wire \ (bool \ bool) spmf" where "xor_protocol A B = do { let (a1, b1) = A; let (a2,b2) = B; return_spmf (a1 \ a2, b1 \ b2)}" lemma lossless_xor_protocol: "lossless_spmf (xor_protocol A B)" by(simp add: xor_protocol_def split_def) lemma share_xor_reconstruct: shows "share x \ (\ w1. share y \ (\ w2. xor_protocol w1 w2 \ (\ (a, b). reconstruct (a, b)))) = xor_evaluate x y" proof- have "(ya = (\ yb)) = ((x = (\ ya)) = (y = (\ yb))) = (x = (\ y))" for ya yb by auto then show ?thesis by(simp add: share_def xor_protocol_def reconstruct_def xor_evaluate_def bind_spmf_const) qed definition R1_xor :: "(bool \ bool) \ (bool \ bool) \ (bool \ bool) spmf" where "R1_xor A B = return_spmf A" lemma lossless_R1_xor: "lossless_spmf (R1_xor A B)" by(simp add: R1_xor_def) definition S1_xor :: "(bool \ bool) \ bool \ (bool \ bool) spmf" where "S1_xor A out = return_spmf A" lemma lossless_S1_xor: "lossless_spmf (S1_xor A out)" by(simp add: S1_xor_def) lemma P1_xor_inf_th: "R1_xor A B = xor_funct A B \ (\ (out1, out2). S1_xor A out1)" by(simp add: R1_xor_def xor_funct_def S1_xor_def split_def) definition R2_xor :: "(bool \ bool) \ (bool \ bool) \ (bool \ bool) spmf" where "R2_xor A B = return_spmf B" lemma lossless_R2_xor: "lossless_spmf (R2_xor A B)" by(simp add: R2_xor_def) definition S2_xor :: "(bool \ bool) \ bool \ (bool \ bool) spmf" where "S2_xor B out = return_spmf B" lemma lossless_S2_xor: "lossless_spmf (S2_xor A out)" by(simp add: S2_xor_def) lemma P2_xor_inf_th: "R2_xor A B = xor_funct A B \ (\ (out1, out2). S2_xor B out2)" by(simp add: R2_xor_def xor_funct_def S2_xor_def split_def) sublocale xor_sim_det: sim_det_def R1_xor S1_xor R2_xor S2_xor xor_funct xor_protocol unfolding sim_det_def_def by(simp add: lossless_R1_xor lossless_S1_xor lossless_R2_xor lossless_S2_xor lossless_xor_funct) -lemma "xor_sim_det.inf_theoretic_P1 m1 m2" - by(simp add: xor_sim_det.inf_theoretic_P1_def P1_xor_inf_th) +lemma "xor_sim_det.perfect_sec_P1 m1 m2" + by(simp add: xor_sim_det.perfect_sec_P1_def P1_xor_inf_th) -lemma "xor_sim_det.inf_theoretic_P2 m1 m2" - by(simp add: xor_sim_det.inf_theoretic_P2_def P2_xor_inf_th) +lemma "xor_sim_det.perfect_sec_P2 m1 m2" + by(simp add: xor_sim_det.perfect_sec_P2_def P2_xor_inf_th) definition and_funct :: "(share_1 \ share_2) \ (share_1 \ share_2) \ share_wire spmf" where "and_funct A B = do { let (a1, a2) = A; let (b1,b2) = B; \ \ coin_spmf; return_spmf (\, \ \ ((a1 \ b1) \ (a2 \ b2)))}" lemma lossless_and_funct: "lossless_spmf (and_funct A B)" by(simp add: and_funct_def split_def) definition and_evaluate :: "bool \ bool \ bool spmf" where "and_evaluate A B = return_spmf (A \ B)" definition and_protocol :: "share_wire \ share_wire \ share_wire spmf" where "and_protocol A B = do { let (a1, b1) = A; let (a2,b2) = B; \ \ coin_spmf; let s0 = \ \ ((a1 \ False) \ (b1 \ False)); let s1 = \ \ ((a1 \ False) \ (b1 \ True)); let s2 = \ \ ((a1 \ True) \ (b1 \ False)); let s3 = \ \ ((a1 \ True) \ (b1 \ True)); (_, s) \ protocol_14_OT (s0,s1,s2,s3) (a2,b2); return_spmf (\, s)}" lemma lossless_and_protocol: "lossless_spmf (and_protocol A B)" by(simp add: and_protocol_def split_def lossless_protocol_14_OT) lemma and_correct: "and_protocol (a1, b1) (a2,b2) = and_funct (a1, b1) (a2,b2)" apply(simp add: and_protocol_def and_funct_def correct_14[symmetric] funct_14) by(cases b2 ; cases b1; cases a1; cases a2; auto) lemma share_and_reconstruct: shows "share x \ (\ (a1,a2). share y \ (\ (b1,b2). and_protocol (a1,b1) (a2,b2) \ (\ (a, b). reconstruct (a, b)))) = and_evaluate x y" proof- have "(yc = (\ (if x = (\ ya) then if snd (snd (ya, x = (\ ya)), snd (yb, y = (\ yb))) then yc = (fst (fst (ya, x = (\ ya)), fst (yb, y = (\ yb))) \ snd (fst (ya, x = (\ ya)), fst (yb, y = (\ yb)))) else yc = (fst (fst (ya, x = (\ ya)), fst (yb, y = (\ yb))) \ \ snd (fst (ya, x = (\ ya)), fst (yb, y = (\ yb)))) else if snd (snd (ya, x = (\ ya)), snd (yb, y = (\ yb))) then yc = (fst (fst (ya, x = (\ ya)), fst (yb, y = (\ yb))) \ snd (fst (ya, x = (\ ya)), fst (yb, y = (\ yb)))) else yc = (fst (fst (ya, x = (\ ya)), fst (yb, y = (\ yb))) \ \ snd (fst (ya, x = (\ ya)), fst (yb, y = (\ yb))))))) = (x \ y)" for yc yb ya by auto then show ?thesis by(auto simp add: share_def reconstruct_def and_protocol_def and_evaluate_def split_def correct_14[symmetric] funct_14 bind_spmf_const Let_def) qed definition and_R1 :: "(share_1 \ share_1) \ (share_2 \ share_2) \ (((share_1 \ share_1) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) spmf" where "and_R1 A B = do { let (a1, a2) = A; let (b1,b2) = B; \ \ coin_spmf; let s0 = \ \ ((a1 \ False) \ (a2 \ False)); let s1 = \ \ ((a1 \ False) \ (a2 \ True)); let s2 = \ \ ((a1 \ True) \ (a2 \ False)); let s3 = \ \ ((a1 \ True) \ (a2 \ True)); V \ R1_14_OT (s0,s1,s2,s3) (b1,b2); (_, s) \ protocol_14_OT (s0,s1,s2,s3) (b1,b2); return_spmf (((a1,a2), \, V), (\, s))}" lemma lossless_and_R1: "lossless_spmf (and_R1 A B)" apply(simp add: and_R1_def split_def lossless_R1_14_OT lossless_protocol_14_OT Let_def) by (metis prod.collapse lossless_R1_14_OT) definition S1_and :: "(share_1 \ share_1) \ bool \ (((bool \ bool) \ bool \ 'v_14_OT1)) spmf" where "S1_and A \ = do { let (a1,a2) = A; let s0 = \ \ ((a1 \ False) \ (a2 \ False)); let s1 = \ \ ((a1 \ False) \ (a2 \ True)); let s2 = \ \ ((a1 \ True) \ (a2 \ False)); let s3 = \ \ ((a1 \ True) \ (a2 \ True)); V \ S1_14_OT (s0,s1,s2,s3) (); return_spmf ((a1,a2), \, V)}" definition out1 :: "(share_1 \ share_1) \ (share_2 \ share_2) \ bool \ (share_1 \ share_2) spmf" where "out1 A B \ = do { let (a1,a2) = A; let (b1,b2) = B; return_spmf (\, \ \ ((a1 \ b1) \ (a2 \ b2)))}" definition S1_and' :: "(share_1 \ share_1) \ (share_2 \ share_2) \ bool \ (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) spmf" where "S1_and' A B \ = do { let (a1,a2) = A; let (b1,b2) = B; let s0 = \ \ ((a1 \ False) \ (a2 \ False)); let s1 = \ \ ((a1 \ False) \ (a2 \ True)); let s2 = \ \ ((a1 \ True) \ (a2 \ False)); let s3 = \ \ ((a1 \ True) \ (a2 \ True)); V \ S1_14_OT (s0,s1,s2,s3) (); return_spmf (((a1,a2), \, V), (\, \ \ ((a1 \ b1) \ (a2 \ b2))))}" lemma sec_ex_P1_and: shows "\ (A :: 'v_14_OT1 \ bool \ bool spmf). \spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True - spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True\ = \spmf (coin_spmf \ (\ \. S1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) () \ (\ view. A view \))) True - spmf (coin_spmf \ (\ \. R1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) (b1, b2) \ (\ view. A view \))) True\" including monad_normalisation proof- define A' where "A' == \ view \. (D (((a1,a2), \, view), (\, \ \ ((a1 \ b1) \ (a2 \ b2)))))" have "\spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True - spmf ((and_R1 (a1, a2) (b1,b2)) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (bool \ bool)) \ bool spmf)) True\ = \spmf (coin_spmf \ (\ \ :: bool. S1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) () \ (\ view. A' view \))) True - spmf (coin_spmf \ (\ \. R1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) (b1, b2) \ (\ view. A' view \))) True\" by(auto simp add: S1_and'_def A'_def and_funct_def and_R1_def Let_def split_def correct_14[symmetric] funct_14; cases a1; cases a2; cases b1; cases b2; auto) then show ?thesis by auto qed lemma bound_14_OT: "\spmf (coin_spmf \ (\ \. S1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) () \ (\ view. (A :: 'v_14_OT1 \ bool \ bool spmf) view \))) True - spmf (coin_spmf \ (\ \. R1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) (b1, b2) \ (\ view. A view \))) True\ \ adv_14_OT" (is "?lhs \ adv_14_OT") proof- have int1: "integrable (measure_spmf coin_spmf) (\x. spmf (S1_14_OT (x \ (a1 \ False \ a2 \ False), x \ (a1 \ False \ a2 \ True), x \ (a1 \ True \ a2 \ False), x \ (a1 \ True \ a2 \ True)) () \ (\view. A view x)) True)" and int2: "integrable (measure_spmf coin_spmf) (\x. spmf (R1_14_OT (x \ (a1 \ False \ a2 \ False), x \ (a1 \ False \ a2 \ True), x \ (a1 \ True \ a2 \ False), x \ (a1 \ True \ a2 \ True)) (b1, b2) \ (\view. A view x)) True)" by(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1)+ have "?lhs = \LINT x|measure_spmf coin_spmf. spmf (S1_14_OT (x \ (a1 \ False \ a2 \ False), x \ (a1 \ False \ a2 \ True), x \ (a1 \ True \ a2 \ False), x \ (a1 \ True \ a2 \ True)) () \ (\view. A view x)) True - spmf (R1_14_OT (x \ (a1 \ False \ a2 \ False), x \ (a1 \ False \ a2 \ True), x \ (a1 \ True \ a2 \ False), x \ (a1 \ True \ a2 \ True)) (b1, b2) \ (\view. A view x)) True\" apply(subst (1 2) spmf_bind) using int1 int2 by simp also have "... \ LINT x|measure_spmf coin_spmf. \spmf (S1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) () \ (\view. A view x)) True - spmf (R1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) (b1, b2) \ (\view. A view x)) True\" by(rule integral_abs_bound[THEN order_trans]; simp add: split_beta) ultimately have "?lhs \ LINT x|measure_spmf coin_spmf. \spmf (S1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) () \ (\view. A view x)) True - spmf (R1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) (b1, b2) \ (\view. A view x)) True\" by simp also have "LINT x|measure_spmf coin_spmf. \spmf (S1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) () \ (\view. A view x)) True - spmf (R1_14_OT (x = (a1 \ \ a2), x = (a1 \ a2), x = (a1 \ \ a2), x = (a1 \ a2)) (b1, b2) \ (\view. A view x)) True\ \ adv_14_OT" apply(rule integral_mono[THEN order_trans]) apply(rule measure_spmf.integrable_const_bound[where B=2]) apply clarsimp apply(rule abs_triangle_ineq4[THEN order_trans]) apply(cases a1) apply(cases a2) subgoal for M using pmf_le_1[of "R1_14_OT (\ M, M, M, M) (b1,b2) \ (\ view. A view M)" "Some True"] pmf_le_1[of "S1_14_OT (\ M, M, M, M) () \ (\ view. A view M)" "Some True"] by simp subgoal for M using pmf_le_1[of "R1_14_OT (M, \ M, M, M) (b1,b2) \ (\ view. A view M)" "Some True"] pmf_le_1[of "S1_14_OT (M, \ M, M, M) () \ (\ view. A view M)" "Some True"] by simp apply(cases a2) apply(auto) subgoal for M using pmf_le_1[of "R1_14_OT (M, M, \ M, M) (b1,b2) \ (\ view. A view M)" "Some True"] pmf_le_1[of "S1_14_OT (M, M, \ M, M) () \ (\ view. A view M)" "Some True"] by(simp) subgoal for M using pmf_le_1[of "R1_14_OT (M, M, M, \ M) (b1,b2) \ (\ view. A view M)" "Some True"] pmf_le_1[of "S1_14_OT (M, M, M, \ M) () \ (\ view. A view M)" "Some True"] by(simp) using ass_adv_14_OT by fast ultimately show ?thesis by simp qed lemma security_and_P1: shows "\spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True - spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True\ \ adv_14_OT" proof- obtain A :: "'v_14_OT1 \ bool \ bool spmf" where A: "\spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ D)) True - spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True\ = \spmf (coin_spmf \ (\ \. S1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) () \ (\ view. A view \))) True - spmf (coin_spmf \ (\ \. R1_14_OT ((\ \ ((a1 \ False) \ (a2 \ False))), (\ \ ((a1 \ False) \ (a2 \ True))), (\ \ ((a1 \ True) \ (a2 \ False))), (\ \ ((a1 \ True) \ (a2 \ True)))) (b1, b2) \ (\ view. A view \))) True\" using sec_ex_P1_and by blast then show ?thesis using bound_14_OT[of "a1" "a2" "A" "b1" "b2" ] by metis qed lemma security_and_P1': shows "\spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True - spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True\ \ adv_14_OT" proof- have "\spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True - spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True\ = \spmf ((and_funct (a1, a2) (b1,b2)) \ (\ (s1, s2). (S1_and' (a1,a2) (b1,b2) s1) \ (D :: (((bool \ bool) \ bool \ 'v_14_OT1) \ (share_1 \ share_2)) \ bool spmf))) True - spmf ((and_R1 (a1, a2) (b1,b2)) \ D) True\" using abs_minus_commute by blast thus ?thesis using security_and_P1 by simp qed definition and_R2 :: "(share_1 \ share_2) \ (share_2 \ share_1) \ (((bool \ bool) \ 'v_14_OT2) \ (share_1 \ share_2)) spmf" where "and_R2 A B = do { let (a1, a2) = A; let (b1,b2) = B; \ \ coin_spmf; let s0 = \ \ ((a1 \ False) \ (a2 \ False)); let s1 = \ \ ((a1 \ False) \ (a2 \ True)); let s2 = \ \ ((a1 \ True) \ (a2 \ False)); let s3 = \ \ ((a1 \ True) \ (a2 \ True)); (_, s) \ protocol_14_OT (s0,s1,s2,s3) B; V \ R2_14_OT (s0,s1,s2,s3) B; return_spmf ((B, V), (\, s))}" lemma lossless_and_R2: "lossless_spmf (and_R2 A B)" apply(simp add: and_R2_def split_def lossless_R2_14_OT lossless_protocol_14_OT Let_def) by (metis lossless_R2_14_OT prod.collapse) definition S2_and :: "(share_1 \ share_2) \ bool \ (((bool \ bool) \ 'v_14_OT2)) spmf" where "S2_and B s2 = do { let (a2,b2) = B; V :: 'v_14_OT2 \ S2_14_OT (a2,b2) s2; return_spmf ((B, V))}" definition out2 :: "(share_1 \ share_2) \ (share_1 \ share_2) \ bool \ (share_1 \ share_2) spmf" where "out2 B A s2 = do { let (a1, b1) = A; let (a2,b2) = B; let s1 = s2 \ ((a1 \ a2) \ (b1 \ b2)); return_spmf (s1, s2)}" definition S2_and' :: "(share_1 \ share_2) \ (share_1 \ share_2) \ bool \ (((bool \ bool) \ 'v_14_OT2) \ (share_1 \ share_2)) spmf" where "S2_and' B A s2 = do { let (a1, a2) = A; let (b1,b2) = B; V :: 'v_14_OT2 \ S2_14_OT B s2; let s1 = s2 \ ((a1 \ b1) \ (a2 \ b2)); return_spmf ((B, V), s1, s2)}" lemma lossless_S2_and: "lossless_spmf (S2_and B s2)" apply(simp add: S2_and_def split_def) by(metis prod.collapse lossless_S2_14_OT) sublocale and_secret_sharing: sim_non_det_def and_R1 S1_and out1 and_R2 S2_and out2 and_funct . lemma ideal_S1_and: "and_secret_sharing.Ideal1 (a1, b1) (a2, b2) s2 = S1_and' (a1, b1) (a2, b2) s2" by(simp add: Let_def and_secret_sharing.Ideal1_def S1_and'_def split_def out1_def S1_and_def) -lemma and_P2_security: "and_secret_sharing.inf_theoretic_P2 m1 m2" +lemma and_P2_security: "and_secret_sharing.perfect_sec_P2 m1 m2" proof- have "and_R2 (a1, b1) (a2, b2) = and_funct (a1, b1) (a2, b2) \ (\(s1, s2). and_secret_sharing.Ideal2 (a2, b2) (a1, b1) s2)" for a1 a2 b1 b2 apply(auto simp add: split_def inf_th_14_OT_P4 S2_and'_def and_R2_def and_funct_def Let_def correct_14[symmetric] and_secret_sharing.Ideal2_def S2_and_def out2_def) apply(simp only: funct_14) apply auto by(cases b1;cases b2; cases a1; cases a2; auto) thus ?thesis - by(simp add: and_secret_sharing.inf_theoretic_P2_def; metis prod.collapse) + by(simp add: and_secret_sharing.perfect_sec_P2_def; metis prod.collapse) qed lemma and_P1_security: "and_secret_sharing.adv_P1 m1 m2 D \ adv_14_OT" proof- have "\spmf (and_R1 (a1, a2) (b1, b2) \ D) True - spmf (and_funct (a1, a2) (b1, b2) \ (\(s1, s2). and_secret_sharing.Ideal1 (a1, a2) (b1, b2) s1 \ D)) True\ \ adv_14_OT" for a1 a2 b1 b2 using security_and_P1' ideal_S1_and prod.collapse by simp thus ?thesis by(simp add: and_secret_sharing.adv_P1_def; metis prod.collapse) qed definition "F = {and_evaluate, xor_evaluate}" lemma share_reconstruct_xor: "share x \ (\(a1, a2). share y \ (\(b1, b2). xor_protocol (a1, b1) (a2, b2) \ (\(a, b). reconstruct (a, b)))) = xor_evaluate x y" proof- have "(((ya = (x = ya)) = (yb = (y = (\ yb))))) = (x = (\ y))" for ya yb by auto thus ?thesis by(simp add: xor_protocol_def share_def reconstruct_def xor_evaluate_def bind_spmf_const) qed sublocale share_correct: secret_sharing_scheme share reconstruct F . lemma "share_correct.sharing_correct input" by(simp add: share_correct.sharing_correct_def reconstruct_share) lemma "share_correct.correct_share_eval input1 input2" unfolding share_correct.correct_share_eval_def apply(auto simp add: F_def) using share_and_reconstruct apply auto using share_reconstruct_xor by force end locale gmw_asym = fixes S1_14_OT :: "nat \ msgs_14_OT \ unit \ 'v_14_OT1 spmf" and R1_14_OT :: "nat \ msgs_14_OT \ choice_14_OT \ 'v_14_OT1 spmf" and S2_14_OT :: "nat \ choice_14_OT \ bool \ 'v_14_OT2 spmf" and R2_14_OT :: "nat \ msgs_14_OT \ choice_14_OT \ 'v_14_OT2 spmf" and protocol_14_OT :: "nat \ msgs_14_OT \ choice_14_OT \ (unit \ bool) spmf" and adv_14_OT :: "nat \ real" assumes gmw_base: "\ (n::nat). gmw_base (S1_14_OT n) (R1_14_OT n) (S2_14_OT n) (R2_14_OT n) (protocol_14_OT n) (adv_14_OT n)" begin sublocale gmw_base "(S1_14_OT n)" "(R1_14_OT n)" "(S2_14_OT n)" "(R2_14_OT n)" "(protocol_14_OT n)" "(adv_14_OT n)" by (simp add: gmw_base) -lemma "xor_sim_det.inf_theoretic_P1 m1 m2" - by (simp add: P1_xor_inf_th xor_sim_det.inf_theoretic_P1_def) +lemma "xor_sim_det.perfect_sec_P1 m1 m2" + by (simp add: P1_xor_inf_th xor_sim_det.perfect_sec_P1_def) -lemma "xor_sim_det.inf_theoretic_P2 m1 m2" - by (simp add: P2_xor_inf_th xor_sim_det.inf_theoretic_P2_def) +lemma "xor_sim_det.perfect_sec_P2 m1 m2" + by (simp add: P2_xor_inf_th xor_sim_det.perfect_sec_P2_def) lemma and_P1_adv_negligible: assumes "negligible (\ n. adv_14_OT n)" shows "negligible (\ n. and_secret_sharing.adv_P1 n m1 m2 D)" proof- have "and_secret_sharing.adv_P1 n m1 m2 D \ adv_14_OT n" for n by (simp add: and_P1_security) thus ?thesis using and_secret_sharing.adv_P1_def assms negligible_le by auto qed -lemma and_P2_security: "and_secret_sharing.inf_theoretic_P2 n m1 m2" +lemma and_P2_security: "and_secret_sharing.perfect_sec_P2 n m1 m2" by (simp add: and_P2_security) end end diff --git a/thys/Multi_Party_Computation/Malicious_Defs.thy b/thys/Multi_Party_Computation/Malicious_Defs.thy --- a/thys/Multi_Party_Computation/Malicious_Defs.thy +++ b/thys/Multi_Party_Computation/Malicious_Defs.thy @@ -1,111 +1,111 @@ section \Malicious Security\ text \Here we define security in the malicious security setting. We follow the definitions given in \cite{DBLP:series/isc/HazayL10} and \cite{DBLP:books/cu/Goldreich2004}. The definition of malicious security follows the real/ideal world paradigm.\ subsection \Malicious Security Definitions\ theory Malicious_Defs imports CryptHOL.CryptHOL begin type_synonym ('in1','aux', 'P1_S1_aux') P1_ideal_adv1 = "'in1' \ 'aux' \ ('in1' \ 'P1_S1_aux') spmf" type_synonym ('in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_ideal_adv2 = "'in1' \ 'aux' \ 'out1' \ 'P1_S1_aux' \ 'adv_out1' spmf" type_synonym ('in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_ideal_adv = "('in1','aux', 'P1_S1_aux') P1_ideal_adv1 \ ('in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_ideal_adv2" type_synonym ('P1_real_adv', 'in1', 'aux', 'P1_S1_aux') P1_sim1 = "'P1_real_adv' \ 'in1' \ 'aux' \ ('in1' \ 'P1_S1_aux') spmf" type_synonym ('P1_real_adv', 'in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_sim2 = "'P1_real_adv' \ 'in1' \ 'aux' \ 'out1' \ 'P1_S1_aux' \ 'adv_out1' spmf" type_synonym ('P1_real_adv', 'in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_sim = "(('P1_real_adv', 'in1', 'aux', 'P1_S1_aux') P1_sim1 \ ('P1_real_adv', 'in1', 'aux', 'out1', 'P1_S1_aux', 'adv_out1') P1_sim2)" type_synonym ('in2','aux', 'P2_S2_aux') P2_ideal_adv1 = "'in2' \ 'aux' \ ('in2' \ 'P2_S2_aux') spmf" type_synonym ('in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_ideal_adv2 = "'in2' \ 'aux' \ 'out2' \ 'P2_S2_aux' \ 'adv_out2' spmf" type_synonym ('in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_ideal_adv = "('in2','aux', 'P2_S2_aux') P2_ideal_adv1 \ ('in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_ideal_adv2" type_synonym ('P2_real_adv', 'in2', 'aux', 'P2_S2_aux') P2_sim1 = "'P2_real_adv' \ 'in2' \ 'aux' \ ('in2' \ 'P2_S2_aux') spmf" type_synonym ('P2_real_adv', 'in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_sim2 = "'P2_real_adv' \ 'in2' \ 'aux' \ 'out2' \ 'P2_S2_aux' \ 'adv_out2' spmf" type_synonym ('P2_real_adv', 'in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_sim = "(('P2_real_adv', 'in2', 'aux', 'P2_S2_aux') P2_sim1 \ ('P2_real_adv', 'in2', 'aux', 'out2', 'P2_S2_aux', 'adv_out2') P2_sim2)" locale malicious_base = fixes funct :: "'in1 \ 'in2 \ ('out1 \ 'out2) spmf" \ \the functionality\ and protocol :: "'in1 \ 'in2 \ ('out1 \ 'out2) spmf" \ \outputs the output of each party in the protocol\ and S1_1 :: "('P1_real_adv, 'in1, 'aux, 'P1_S1_aux) P1_sim1" \ \first part of the simulator for party 1\ and S1_2 :: "('P1_real_adv, 'in1, 'aux, 'out1, 'P1_S1_aux, 'adv_out1) P1_sim2" \ \second part of the simulator for party 1\ and P1_real_view :: "'in1 \ 'in2 \ 'aux \ 'P1_real_adv \ ('adv_out1 \ 'out2) spmf" \ \real view for party 1, the adversary totally controls party 1\ and S2_1 :: "('P2_real_adv, 'in2, 'aux, 'P2_S2_aux) P2_sim1" \ \first part of the simulator for party 2\ and S2_2 :: "('P2_real_adv, 'in2, 'aux, 'out2, 'P2_S2_aux, 'adv_out2) P2_sim2" \ \second part of the simulator for party 1\ and P2_real_view :: "'in1 \ 'in2 \ 'aux \ 'P2_real_adv \ ('out1 \ 'adv_out2) spmf" \ \real view for party 2, the adversary totally controls party 2\ begin definition "correct m1 m2 \ (protocol m1 m2 = funct m1 m2)" abbreviation "trusted_party x y \ funct x y" text\The ideal game defines the ideal world. First we consider the case where party 1 is corrupt, and thus controlled by the adversary. The adversary is split into two parts, the first part takes the original input and auxillary information and outputs its input to the protocol. The trusted party then computes the functionality on the input given by the adversary and the correct input for party 2. Party 2 outputs the its correct output as given by the trusted party, the adversary provides the output for party 1.\ definition ideal_game_1 :: "'in1 \ 'in2 \ 'aux \ ('in1, 'aux, 'out1, 'P1_S1_aux, 'adv_out1) P1_ideal_adv \ ('adv_out1 \ 'out2) spmf" where "ideal_game_1 x y z A = do { let (A1,A2) = A; (x', aux_out) \ A1 x z; (out1, out2) \ trusted_party x' y; out1' :: 'adv_out1 \ A2 x' z out1 aux_out; return_spmf (out1', out2)}" definition ideal_view_1 :: "'in1 \ 'in2 \ 'aux \ ('P1_real_adv, 'in1, 'aux, 'out1, 'P1_S1_aux, 'adv_out1) P1_sim \ 'P1_real_adv \ ('adv_out1 \ 'out2) spmf" where "ideal_view_1 x y z S \ = (let (S1, S2) = S in (ideal_game_1 x y z (S1 \, S2 \)))" text\We have information theoretic security when the real and ideal views are equal.\ -definition "inf_theoretic_mal_sec_P1 x y z S \ \ (ideal_view_1 x y z S \ = P1_real_view x y z \)" +definition "perfect_sec_P1 x y z S \ \ (ideal_view_1 x y z S \ = P1_real_view x y z \)" text\The advantage of party 1 denotes the probability of a distinguisher distinguishing the real and ideal views.\ definition "adv_P1 x y z S \ (D :: ('adv_out1 \ 'out2) \ bool spmf) = \spmf (P1_real_view x y z \ \ (\ view. D view)) True - spmf (ideal_view_1 x y z S \ \ (\ view. D view)) True \" definition ideal_game_2 :: "'in1 \ 'in2 \ 'aux \ ('in2, 'aux, 'out2, 'P2_S2_aux, 'adv_out2) P2_ideal_adv \ ('out1 \ 'adv_out2) spmf" where "ideal_game_2 x y z A = do { let (A1,A2) = A; (y', aux_out) \ A1 y z; (out1, out2) \ trusted_party x y'; out2' :: 'adv_out2 \ A2 y' z out2 aux_out; return_spmf (out1, out2')}" definition ideal_view_2 :: "'in1 \ 'in2 \ 'aux \ ('P2_real_adv, 'in2, 'aux, 'out2, 'P2_S2_aux, 'adv_out2) P2_sim \ 'P2_real_adv \ ('out1 \ 'adv_out2) spmf" where "ideal_view_2 x y z S \ = (let (S1, S2) = S in (ideal_game_2 x y z (S1 \, S2 \)))" -definition "inf_theoretic_mal_sec_P2 x y z S \ \ (ideal_view_2 x y z S \ = P2_real_view x y z \)" +definition "perfect_sec_P2 x y z S \ \ (ideal_view_2 x y z S \ = P2_real_view x y z \)" definition "adv_P2 x y z S \ (D :: ('out1 \ 'adv_out2) \ bool spmf) = \spmf (P2_real_view x y z \ \ (\ view. D view)) True - spmf (ideal_view_2 x y z S \ \ (\ view. D view)) True \" end end \ No newline at end of file diff --git a/thys/Multi_Party_Computation/Malicious_OT.thy b/thys/Multi_Party_Computation/Malicious_OT.thy --- a/thys/Multi_Party_Computation/Malicious_OT.thy +++ b/thys/Multi_Party_Computation/Malicious_OT.thy @@ -1,1542 +1,1542 @@ subsection \Malicious OT\ text \Here we prove secure the 1-out-of-2 OT protocol given in \cite{DBLP:series/isc/HazayL10} (p190). For party 1 reduce security to the DDH assumption and for party 2 we show information theoretic security.\ theory Malicious_OT imports "HOL-Number_Theory.Cong" Cyclic_Group_Ext DH_Ext Malicious_Defs Number_Theory_Aux OT_Functionalities Uniform_Sampling begin type_synonym ('aux, 'grp', 'state) adv_1_P1 = "('grp' \ 'grp') \ 'grp' \ 'grp' \ 'grp' \ 'grp' \ 'grp' \ 'aux \ (('grp' \ 'grp' \ 'grp') \ 'state) spmf" type_synonym ('grp', 'state) adv_2_P1 = "'grp' \ 'grp' \ 'grp' \ 'grp' \ 'grp' \ ('grp' \ 'grp') \ 'state \ ((('grp' \ 'grp') \ ('grp' \ 'grp')) \ 'state) spmf" type_synonym ('adv_out1,'state) adv_3_P1 = "'state \ 'adv_out1 spmf" type_synonym ('aux, 'grp', 'adv_out1, 'state) adv_mal_P1 = "(('aux, 'grp', 'state) adv_1_P1 \ ('grp', 'state) adv_2_P1 \ ('adv_out1,'state) adv_3_P1)" type_synonym ('aux, 'grp','state) adv_1_P2 = "bool \ 'aux \ (('grp' \ 'grp' \ 'grp' \ 'grp' \ 'grp') \ 'state) spmf" type_synonym ('grp','state) adv_2_P2 = "('grp' \ 'grp' \ 'grp' \ 'grp' \ 'grp') \ 'state \ ((('grp' \ 'grp' \ 'grp') \ nat) \ 'state) spmf" type_synonym ('grp', 'adv_out2, 'state) adv_3_P2 = "('grp' \ 'grp') \ ('grp' \ 'grp') \ 'state \ 'adv_out2 spmf" type_synonym ('aux, 'grp', 'adv_out2, 'state) adv_mal_P2 = "(('aux, 'grp','state) adv_1_P2 \ ('grp','state) adv_2_P2 \ ('grp', 'adv_out2,'state) adv_3_P2)" locale ot_base = fixes \ :: "'grp cyclic_group" (structure) assumes finite_group: "finite (carrier \)" and order_gt_0: "order \ > 0" and prime_order: "prime (order \)" begin lemma prime_field: "a < (order \) \ a \ 0 \ coprime a (order \)" by (metis dvd_imp_le neq0_conv not_le prime_imp_coprime prime_order coprime_commute) text\The protocol uses a call to an idealised functionality of a zero knowledge protocol for the DDH relation, this is described by the functionality given below.\ fun funct_DH_ZK :: "('grp \ 'grp \ 'grp) \ (('grp \ 'grp \ 'grp) \ nat) \ (bool \ unit) spmf" where "funct_DH_ZK (h,a,b) ((h',a',b'),r) = return_spmf (a = \<^bold>g [^] r \ b = h [^] r \ (h,a,b) = (h',a',b'), ())" text\The probabilistic program that defines the output for both parties in the protocol.\ definition protocol_ot :: "('grp \ 'grp) \ bool \ (unit \ 'grp) spmf" where "protocol_ot M \ = do { let (x0,x1) = M; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = h0 [^] r \ \<^bold>g [^] (if \ then (1::nat) else 0); let b1 = h1 [^] r \ \<^bold>g [^] (if \ then (1::nat) else 0); let h = h0 \ inv h1; let b = b0 \ inv b1; _ :: unit \ assert_spmf (a = \<^bold>g [^] r \ b = h [^] r); u0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v0 \ sample_uniform (order \); v1 \ sample_uniform (order \); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let w0 = a [^] u0 \ \<^bold>g [^] v0; let e0 = (w0, z0); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let w1 = a [^] u1 \ \<^bold>g [^] v1; let e1 = (w1, z1); return_spmf ((), (if \ then (z1 \ inv (w1 [^] \1)) else (z0 \ inv (w0 [^] \0))))}" text\Party 1 sends three messages (including the output) in the protocol so we split the adversary into three parts, one part to output each message. The real view of the protocol for party 1 outputs the correct output for party 2 and the adversary outputs the output for party 1.\ definition P1_real_model :: "('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out1, 'state) adv_mal_P1 \ ('adv_out1 \ 'grp) spmf" where "P1_real_model M \ z \ = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = h0 [^] r \ (if \ then \<^bold>g else \); let b1 = h1 [^] r \ (if \ then \<^bold>g else \); ((in1 :: 'grp, in2 ::'grp, in3 :: 'grp), s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (b :: bool, _ :: unit) \ funct_DH_ZK (in1, in2, in3) ((h,a,b), r); _ :: unit \ assert_spmf (b); (((w0,z0),(w1,z1)), s') \ \2 h0 h1 a b0 b1 M s; adv_out :: 'adv_out1 \ \3 s'; return_spmf (adv_out, (if \ then (z1 \ (inv w1 [^] \1)) else (z0 \ (inv w0 [^] \0))))}" text\The first and second part of the simulator for party 1 are defined below.\ definition P1_S1 :: "('aux, 'grp, 'adv_out1, 'state) adv_mal_P1 \ ('grp \ 'grp) \ 'aux \ (('grp \ 'grp) \ 'state) spmf" where "P1_S1 \ M z = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = h0 [^] r; let b1 = h1 [^] r \ \<^bold>g; ((in1 :: 'grp, in2 ::'grp, in3 :: 'grp), s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); _ :: unit \ assert_spmf ((in1, in2, in3) = (h,a,b)); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x0 = (z0 \ (inv w0 [^] \0)); let x1 = (z1 \ (inv w1 [^] \1)); return_spmf ((x0,x1), s')}" definition P1_S2 :: "('aux, 'grp, 'adv_out1,'state) adv_mal_P1 \ ('grp \ 'grp) \ 'aux \ unit \ 'state \ 'adv_out1 spmf" where "P1_S2 \ M z out1 s' = do { let (\1, \2, \3) = \; \3 s'}" text\We explicitly provide the unfolded definition of the ideal model for convieience in the proof.\ definition P1_ideal_model :: "('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out1,'state) adv_mal_P1 \ ('adv_out1 \ 'grp) spmf" where "P1_ideal_model M \ z \ = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = h0 [^] r; let b1 = h1 [^] r \ \<^bold>g; ((in1 :: 'grp, in2 ::'grp, in3 :: 'grp), s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); _ :: unit \ assert_spmf ((in1, in2, in3) = (h,a,b)); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x0' = z0 \ inv w0 [^] \0; let x1' = z1 \ inv w1 [^] \1; (_, f_out2) \ funct_OT_12 (x0', x1') \; adv_out :: 'adv_out1 \ \3 s'; return_spmf (adv_out, f_out2)}" text\The advantage associated with the unfolded definition of the ideal view.\ definition "P1_adv_real_ideal_model (D :: ('adv_out1 \ 'grp) \ bool spmf) M \ \ z = \spmf ((P1_real_model M \ z \) \ (\ view. D view)) True - spmf ((P1_ideal_model M \ z \) \ (\ view. D view)) True\" text\We now define the real view and simulators for party 2 in an analogous way.\ definition P2_real_model :: "('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out2,'state) adv_mal_P2 \ (unit \ 'adv_out2) spmf" where "P2_real_model M \ z \ = do { let (x0,x1) = M; let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3 :: 'grp), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; u0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v0 \ sample_uniform (order \); v1 \ sample_uniform (order \); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let w0 = a [^] u0 \ \<^bold>g [^] v0; let e0 = (w0, z0); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let w1 = a [^] u1 \ \<^bold>g [^] v1; let e1 = (w1, z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" definition P2_S1 :: "('aux, 'grp, 'adv_out2,'state) adv_mal_P2 \ bool \ 'aux \ (bool \ ('grp \ 'grp \ 'grp \ 'grp \ 'grp) \ 'state) spmf" where "P2_S1 \ \ z = do { let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3 :: 'grp), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; let l = b0 \ (inv (h0 [^] r)); return_spmf ((if l = \ then False else True), (h0,h1,a,b0,b1), s')}" definition P2_S2 :: "('aux, 'grp, 'adv_out2,'state) adv_mal_P2 \ bool \ 'aux \ 'grp \ (('grp \ 'grp \ 'grp \ 'grp \ 'grp) \ 'state) \ 'adv_out2 spmf" where "P2_S2 \ \' z x\ aux_out = do { let (\1, \2, \3) = \; let ((h0,h1,a,b0,b1),s) = aux_out; u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = a [^] u0 \ \<^bold>g [^] v0; let w1 = a [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ (if \' then \ else x\); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ (if \' then x\ else \); let e0 = (w0,z0); let e1 = (w1,z1); \3 e0 e1 s}" sublocale mal_def : malicious_base funct_OT_12 protocol_ot P1_S1 P1_S2 P1_real_model P2_S1 P2_S2 P2_real_model . text\We prove the unfolded definition of the ideal views are equal to the definition we provide in the abstract locale that defines security.\ lemma P1_ideal_ideal_eq: shows "mal_def.ideal_view_1 x y z (P1_S1, P1_S2) \ = P1_ideal_model x y z \" including monad_normalisation unfolding mal_def.ideal_view_1_def mal_def.ideal_game_1_def P1_ideal_model_def P1_S1_def P1_S2_def Let_def split_def by(simp add: split_def) lemma P1_advantages_eq: shows "mal_def.adv_P1 x y z (P1_S1, P1_S2) \ D = P1_adv_real_ideal_model D x y \ z" by(simp add: mal_def.adv_P1_def P1_adv_real_ideal_model_def P1_ideal_ideal_eq) fun P1_DDH_mal_adv_\_false :: "('grp \ 'grp) \ 'aux \ ('aux, 'grp, 'adv_out1,'state) adv_mal_P1 \ (('adv_out1 \ 'grp) \ bool spmf) \ 'grp ddh.adversary" where "P1_DDH_mal_adv_\_false M z \ D h a t = do { let (\1, \2, \3) = \; \0 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = h; let b0 = a [^] \0; let b1 = t; ((in1 :: 'grp, in2 ::'grp, in3 :: 'grp),s) \ \1 M h0 h1 a b0 b1 z; _ :: unit \ assert_spmf (in1 = h0 \ inv h1 \ in2 = a \ in3 = b0 \ inv b1); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x0 = (z0 \ (inv w0 [^] \0)); adv_out :: 'adv_out1 \ \3 s'; D (adv_out, x0)}" fun P1_DDH_mal_adv_\_true :: "('grp \ 'grp) \ 'aux \ ('aux, 'grp, 'adv_out1,'state) adv_mal_P1 \ (('adv_out1 \ 'grp) \ bool spmf) \ 'grp ddh.adversary" where "P1_DDH_mal_adv_\_true M z \ D h a t = do { let (\1, \2, \3) = \; \1 :: nat \ sample_uniform (order \); let h1 = \<^bold>g [^] \1; let h0 = h; let b0 = t; let b1 = a [^] \1 \ \<^bold>g; ((in1 :: 'grp, in2 ::'grp, in3 :: 'grp),s) \ \1 M h0 h1 a b0 b1 z; _ :: unit \ assert_spmf (in1 = h0 \ inv h1 \ in2 = a \ in3 = b0 \ inv b1); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x1 = (z1 \ (inv w1 [^] \1)); adv_out :: 'adv_out1 \ \3 s'; D (adv_out, x1)}" definition P2_ideal_model :: "('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out2, 'state) adv_mal_P2 \ (unit \ 'adv_out2) spmf" where "P2_ideal_model M \ z \ = do { let (x0,x1) = M; let (\1, \2, \3) = \; ((h0,h1,a,b0,b1), s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; let l = b0 \ (inv (h0 [^] r)); let \' = (if l = \ then False else True); (_ :: unit, x\) \ funct_OT_12 (x0, x1) \'; u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = a [^] u0 \ \<^bold>g [^] v0; let w1 = a [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ (if \' then \ else x\); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ (if \' then x\ else \); let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" definition P2_ideal_model_end :: "('grp \ 'grp) \ 'grp \ (('grp \ 'grp \ 'grp \ 'grp \ 'grp) \ 'state) \ ('grp, 'adv_out2, 'state) adv_3_P2 \ (unit \ 'adv_out2) spmf" where "P2_ideal_model_end M l bs \3 = do { let (x0,x1) = M; let ((h0,h1,a,b0,b1),s) = bs; let \' = (if l = \ then False else True); (_:: unit, x\) \ funct_OT_12 (x0, x1) \'; u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = a [^] u0 \ \<^bold>g [^] v0; let w1 = a [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ (if \' then \ else x\); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ (if \' then x\ else \); let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s; return_spmf ((), out)}" definition P2_ideal_model' :: "('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out2, 'state) adv_mal_P2 \ (unit \ 'adv_out2) spmf" where "P2_ideal_model' M \ z \ = do { let (x0,x1) = M; let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3 :: 'grp), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; let l = b0 \ (inv (h0 [^] r)); P2_ideal_model_end (x0,x1) l ((h0,h1,a,b0,b1),s') \3}" lemma P2_ideal_model_rewrite: "P2_ideal_model M \ z \ = P2_ideal_model' M \ z \ " by(simp add: P2_ideal_model'_def P2_ideal_model_def P2_ideal_model_end_def Let_def split_def) definition P2_real_model_end :: "('grp \ 'grp) \ (('grp \ 'grp \ 'grp \ 'grp \ 'grp) \ 'state) \ ('grp, 'adv_out2,'state) adv_3_P2 \ (unit \ 'adv_out2) spmf" where "P2_real_model_end M bs \3 = do { let (x0,x1) = M; let ((h0,h1,a,b0,b1),s) = bs; u0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v0 \ sample_uniform (order \); v1 \ sample_uniform (order \); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let w0 = a [^] u0 \ \<^bold>g [^] v0; let e0 = (w0, z0); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let w1 = a [^] u1 \ \<^bold>g [^] v1; let e1 = (w1, z1); out \ \3 e0 e1 s; return_spmf ((), out)}" definition P2_real_model' ::"('grp \ 'grp) \ bool \ 'aux \ ('aux, 'grp, 'adv_out2, 'state) adv_mal_P2 \ (unit \ 'adv_out2) spmf" where "P2_real_model' M \ z \ = do { let (x0,x1) = M; let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3 :: 'grp), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; P2_real_model_end M ((h0,h1,a,b0,b1),s') \3}" lemma P2_real_model_rewrite: "P2_real_model M \ z \ = P2_real_model' M \ z \" by(simp add: P2_real_model'_def P2_real_model_def P2_real_model_end_def split_def) lemma P2_ideal_view_unfold: "mal_def.ideal_view_2 (x0,x1) \ z (P2_S1, P2_S2) \ = P2_ideal_model (x0,x1) \ z \" unfolding local.mal_def.ideal_view_2_def P2_ideal_model_def local.mal_def.ideal_game_2_def P2_S1_def P2_S2_def by(auto simp add: Let_def split_def) end locale ot = ot_base + cyclic_group \ begin lemma P1_assert_correct1: shows "((\<^bold>g [^] (\0::nat)) [^] (r::nat) \ \<^bold>g \ inv ((\<^bold>g [^] (\1::nat)) [^] r \ \<^bold>g) = (\<^bold>g [^] \0 \ inv (\<^bold>g [^] \1)) [^] r)" (is "?lhs = ?rhs") proof- have in_carrier1: "(\<^bold>g [^] \1) [^] r \ carrier \" by simp have in_carrier2: "inv ((\<^bold>g [^] \1) [^] r) \ carrier \" by simp have 1:"?lhs = (\<^bold>g [^] \0) [^] r \ \<^bold>g \ inv ((\<^bold>g [^] \1) [^] r) \ inv \<^bold>g" using cyclic_group_assoc nat_pow_pow inverse_split in_carrier1 by simp also have 2:"... = (\<^bold>g [^] \0) [^] r \ (\<^bold>g \ inv ((\<^bold>g [^] \1) [^] r)) \ inv \<^bold>g" using cyclic_group_assoc in_carrier2 by simp also have "... = (\<^bold>g [^] \0) [^] r \ (inv ((\<^bold>g [^] \1) [^] r) \ \<^bold>g) \ inv \<^bold>g" using in_carrier2 cyclic_group_commute by simp also have 3: "... = (\<^bold>g [^] \0) [^] r \ inv ((\<^bold>g [^] \1) [^] r) \ (\<^bold>g \ inv \<^bold>g)" using cyclic_group_assoc in_carrier2 by simp also have "... = (\<^bold>g [^] \0) [^] r \ inv ((\<^bold>g [^] \1) [^] r)" by simp also have "... = (\<^bold>g [^] \0) [^] r \ inv ((\<^bold>g [^] \1)) [^] r" using inverse_pow_pow by simp ultimately show ?thesis by (simp add: cyclic_group_commute pow_mult_distrib) qed lemma P1_assert_correct2: shows "(\<^bold>g [^] (\0::nat)) [^] (r::nat) \ inv ((\<^bold>g [^] (\1::nat)) [^] r) = (\<^bold>g [^] \0 \ inv (\<^bold>g [^] \1)) [^] r" (is "?lhs = ?rhs") proof- have in_carrier2:"\<^bold>g [^] \1 \ carrier \" by simp hence "?lhs = (\<^bold>g [^] \0) [^] r \ inv ((\<^bold>g [^] \1)) [^] r" using inverse_pow_pow by simp thus ?thesis by (simp add: cyclic_group_commute monoid_comm_monoidI pow_mult_distrib) qed sublocale ddh: ddh_ext by (simp add: cyclic_group_axioms ddh_ext.intro) lemma P1_real_ddh0_\_false: assumes "\ = False" shows "((P1_real_model M \ z \) \ (\ view. D view)) = (ddh.DDH0 (P1_DDH_mal_adv_\_false M z \ D))" including monad_normalisation proof- have "(in2 = \<^bold>g [^] (r::nat) \ in3 = in1 [^] r \ in1 = \<^bold>g [^] (\0::nat) \ inv (\<^bold>g [^] (\1::nat)) \ in2 = \<^bold>g [^] r \ in3 = (\<^bold>g [^] r) [^] \0 \ inv ((\<^bold>g [^] \1) [^] r)) \ (in1 = \<^bold>g [^] \0 \ inv (\<^bold>g [^] \1) \ in2 = \<^bold>g [^] r \ in3 = (\<^bold>g [^] r) [^] \0 \ inv ((\<^bold>g [^] \1) [^] r))" for in1 in2 in3 r \0 \1 by (auto simp add: P1_assert_correct2 power_swap) moreover have "((P1_real_model M \ z \) \ (\ view. D view)) = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = (\<^bold>g [^] r) [^] \0; let b1 = h1 [^] r; ((in1, in2, in3),s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (b :: bool, _ :: unit) \ funct_DH_ZK (in1, in2, in3) ((h,a,b), r); _ :: unit \ assert_spmf (b); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; adv_out \ \3 s'; D (adv_out, ((z0 \ (inv w0 [^] \0))))}" by(simp add: P1_real_model_def assms split_def Let_def power_swap) ultimately show ?thesis by(simp add: P1_real_model_def ddh.DDH0_def Let_def) qed lemma P1_ideal_ddh1_\_false: assumes "\ = False" shows "((P1_ideal_model M \ z \) \ (\ view. D view)) = (ddh.DDH1 (P1_DDH_mal_adv_\_false M z \ D))" including monad_normalisation proof- have "((P1_ideal_model M \ z \) \ (\ view. D view)) = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = (\<^bold>g [^] r) [^] \0; let b1 = h1 [^] r \ \<^bold>g; ((in1, in2, in3),s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); _ :: unit \ assert_spmf ((in1, in2, in3) = (h,a,b)); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x0 = (z0 \ (inv w0 [^] \0)); let x1 = (z1 \ (inv w1 [^] \1)); (_, f_out2) \ funct_OT_12 (x0, x1) \; adv_out \ \3 s'; D (adv_out, f_out2)}" by(simp add: P1_ideal_model_def assms split_def Let_def power_swap) thus ?thesis by(auto simp add: P1_ideal_model_def ddh.DDH1_def funct_OT_12_def Let_def assms ) qed lemma P1_real_ddh1_\_true: assumes "\ = True" shows "((P1_real_model M \ z \) \ (\ view. D view)) = (ddh.DDH1 (P1_DDH_mal_adv_\_true M z \ D))" including monad_normalisation proof- have "(in2 = \<^bold>g [^] (r::nat) \ in3 = in1 [^] r \ in1 = \<^bold>g [^] (\0::nat) \ inv (\<^bold>g [^] (\1::nat)) \ in2 = \<^bold>g [^] r \ in3 = (\<^bold>g [^] r) [^] \0 \ \<^bold>g \ inv ((\<^bold>g [^] \1) [^] r \ \<^bold>g)) \ (in1 = \<^bold>g [^] \0 \ inv (\<^bold>g [^] \1) \ in2 = \<^bold>g [^] r \ in3 = (\<^bold>g [^] \0) [^] r \ \<^bold>g \ inv ((\<^bold>g [^] r) [^] \1 \ \<^bold>g))" for in1 in2 in3 r \0 \1 by (auto simp add: P1_assert_correct1 power_swap) moreover have "((P1_real_model M \ z \) \ (\ view. D view)) = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = ((\<^bold>g [^] r) [^] \0) \ \<^bold>g; let b1 = (h1 [^] r) \ \<^bold>g; ((in1, in2, in3),s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (b :: bool, _ :: unit) \ funct_DH_ZK (in1, in2, in3) ((h,a,b), r); _ :: unit \ assert_spmf (b); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; adv_out \ \3 s'; D (adv_out, ((z1 \ (inv w1 [^] \1))))}" by(simp add: P1_real_model_def assms split_def Let_def power_swap) ultimately show ?thesis by(simp add: Let_def P1_real_model_def ddh.DDH1_def assms power_swap) qed lemma P1_ideal_ddh0_\_true: assumes "\ = True" shows "((P1_ideal_model M \ z \) \ (\ view. D view)) = (ddh.DDH0 (P1_DDH_mal_adv_\_true M z \ D))" including monad_normalisation proof- have "((P1_ideal_model M \ z \) \ (\ view. D view)) = do { let (\1, \2, \3) = \; r \ sample_uniform (order \); \0 \ sample_uniform (order \); \1 \ sample_uniform (order \); let h0 = \<^bold>g [^] \0; let h1 = \<^bold>g [^] \1; let a = \<^bold>g [^] r; let b0 = (\<^bold>g [^] r) [^] \0; let b1 = h1 [^] r \ \<^bold>g; ((in1, in2, in3),s) \ \1 M h0 h1 a b0 b1 z; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); _ :: unit \ assert_spmf ((in1, in2, in3) = (h,a,b)); (((w0,z0),(w1,z1)),s') \ \2 h0 h1 a b0 b1 M s; let x0 = (z0 \ (inv w0 [^] \0)); let x1 = (z1 \ (inv w1 [^] \1)); (_, f_out2) \ funct_OT_12 (x0, x1) \; adv_out \ \3 s'; D (adv_out, f_out2)}" by(simp add: P1_ideal_model_def assms Let_def split_def power_swap) thus ?thesis by(simp add: split_def Let_def P1_ideal_model_def ddh.DDH0_def assms funct_OT_12_def power_swap) qed lemma P1_real_ideal_DDH_advantage_false: assumes "\ = False" shows "mal_def.adv_P1 M \ z (P1_S1, P1_S2) \ D = ddh.DDH_advantage (P1_DDH_mal_adv_\_false M z \ D)" by(simp add: P1_adv_real_ideal_model_def ddh.DDH_advantage_def P1_ideal_ddh1_\_false P1_real_ddh0_\_false assms P1_advantages_eq) lemma P1_real_ideal_DDH_advantage_false_bound: assumes "\ = False" shows "mal_def.adv_P1 M \ z (P1_S1, P1_S2) \ D \ ddh.advantage (P1_DDH_mal_adv_\_false M z \ D) + ddh.advantage (ddh.DDH_\' (P1_DDH_mal_adv_\_false M z \ D))" using P1_real_ideal_DDH_advantage_false ddh.DDH_advantage_bound assms by metis lemma P1_real_ideal_DDH_advantage_true: assumes "\ = True" shows "mal_def.adv_P1 M \ z (P1_S1, P1_S2) \ D = ddh.DDH_advantage (P1_DDH_mal_adv_\_true M z \ D)" by(simp add: P1_adv_real_ideal_model_def ddh.DDH_advantage_def P1_real_ddh1_\_true P1_ideal_ddh0_\_true assms P1_advantages_eq) lemma P1_real_ideal_DDH_advantage_true_bound: assumes "\ = True" shows "mal_def.adv_P1 M \ z (P1_S1, P1_S2) \ D \ ddh.advantage (P1_DDH_mal_adv_\_true M z \ D) + ddh.advantage (ddh.DDH_\' (P1_DDH_mal_adv_\_true M z \ D))" using P1_real_ideal_DDH_advantage_true ddh.DDH_advantage_bound assms by metis (*Auxillary proofs that we use in the proof of security, mainly rewriting things for P2*) lemma P2_output_rewrite: assumes "s < order \" shows "(\<^bold>g [^] (r * u1 + v1), \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] u1) = (\<^bold>g [^] (r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \), \<^bold>g [^] (r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) \ inv \<^bold>g [^] ((s + u1) mod order \ + (order \ - s)))" proof- have "\<^bold>g [^] (r * u1 + v1) = \<^bold>g [^] (r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \)" proof- have "[(r * u1 + v1) = (r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \)] (mod (order \))" proof- have "[(r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \) = r * (s + u1) + (r * order \ - r * s + v1)] (mod (order \))" by (metis (no_types, hide_lams) cong_def mod_add_left_eq mod_add_right_eq mod_mult_right_eq) hence "[(r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \) = r * s + r * u1 + r * order \ - r * s + v1] (mod (order \))" by (metis (no_types, lifting) Nat.add_diff_assoc add.assoc assms distrib_left less_or_eq_imp_le mult_le_mono) hence "[(r * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \) = r * u1 + r * order \ + v1] (mod (order \))" by simp thus ?thesis by (simp add: cong_def semiring_normalization_rules(23)) qed then show ?thesis using finite_group pow_generator_eq_iff_cong by blast qed moreover have " \<^bold>g [^] (r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) \ inv \<^bold>g [^] ((s + u1) mod order \ + (order \ - s)) = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] u1" proof- have "\<^bold>g [^] (r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = \<^bold>g [^] (r * \ * u1 + v1 * \)" proof- have "[(r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = r * \ * u1 + v1 * \] (mod (order \))" proof- have "[(r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = r * \ * (s + u1) + (r * order \ - r * s + v1) * \] (mod (order \))" using cong_def mod_add_cong mod_mult_left_eq mod_mult_right_eq by blast hence "[(r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = r * \ * s + r * \ * u1 + (r * order \ - r * s + v1) * \] (mod (order \))" by (simp add: distrib_left) hence "[(r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = r * \ * s + r * \ * u1 + r * order \ * \ - r * s * \ + v1 * \] (mod (order \))" using distrib_right assms by (smt Groups.mult_ac(3) order_gt_0 Nat.add_diff_assoc2 add.commute diff_mult_distrib2 mult.commute mult_strict_mono order.strict_implies_order semiring_normalization_rules(25) zero_order(1)) hence "[(r * \ * ((s + u1) mod order \) + (r * order \ - r * s + v1) mod order \ * \) = r * \ * u1 + r * order \ * \ + v1 * \] (mod (order \))" by simp thus ?thesis by (simp add: cong_def semiring_normalization_rules(16) semiring_normalization_rules(23)) qed thus ?thesis using finite_group pow_generator_eq_iff_cong by blast qed also have "inv \<^bold>g [^] ((s + u1) mod order \ + (order \ - s)) = inv \<^bold>g [^] u1" proof- have "[((s + u1) mod order \ + (order \ - s)) = u1] (mod (order \))" proof- have "[((s + u1) mod order \ + (order \ - s)) = s + u1 + (order \ - s)] (mod (order \))" by (simp add: add.commute mod_add_right_eq cong_def) hence "[((s + u1) mod order \ + (order \ - s)) = u1 + order \] (mod (order \))" using assms by simp thus ?thesis by (simp add: cong_def) qed hence "\<^bold>g [^] ((s + u1) mod order \ + (order \ - s)) = \<^bold>g [^] u1" using finite_group pow_generator_eq_iff_cong by blast thus ?thesis by (metis generator_closed inverse_pow_pow) qed ultimately show ?thesis by argo qed ultimately show ?thesis by simp qed lemma P2_inv_g_rewrite: assumes "s < order \" shows "(inv \<^bold>g) [^] (u1' + (order \ - s)) = \<^bold>g [^] s \ inv (\<^bold>g [^] u1')" proof- have power_commute_rewrite: "\<^bold>g [^] (((order \ - s) + u1') mod order \) = \<^bold>g [^] (u1' + (order \ - s))" using add.commute pow_generator_mod by metis have "(order \ - s + u1') mod order \ = (int (order \) - int s + int u1') mod order \" by (metis of_nat_add of_nat_diff order.strict_implies_order zmod_int assms(1)) also have "... = (- int s + int u1') mod order \" by (metis (full_types) add.commute minus_mod_self1 mod_add_right_eq) ultimately have "(order \ - s + u1') mod order \ = (- int s mod (order \) + int u1' mod (order \)) mod order \" by presburger hence "\<^bold>g [^] (((order \ - s) + u1') mod order \) = \<^bold>g [^] ((- int s mod (order \) + int u1' mod (order \)) mod order \)" by (metis int_pow_int) hence "\<^bold>g [^] (u1' + (order \ - s)) = \<^bold>g [^] ((- int s mod (order \) + int u1' mod (order \)) mod order \)" using power_commute_rewrite by argo also have "... = \<^bold>g [^] (- int s mod (order \) + int u1' mod (order \))" using pow_generator_mod_int by blast also have "... = \<^bold>g [^] (- int s mod (order \)) \ \<^bold>g [^] (int u1' mod (order \))" by (simp add: int_pow_mult) also have "... = \<^bold>g [^] (- int s) \ \<^bold>g [^] (int u1')" using pow_generator_mod_int by simp ultimately have "inv (\<^bold>g [^] (u1' + (order \ - s))) = inv (\<^bold>g [^] (- int s) \ \<^bold>g [^] (int u1'))" by simp hence "inv (\<^bold>g [^] ((u1' + (order \ - s)) mod (order \)) ) = inv (\<^bold>g [^] (- int s)) \ inv (\<^bold>g [^] (int u1'))" using pow_generator_mod by (simp add: inverse_split) also have "... = \<^bold>g [^] (int s) \ inv (\<^bold>g [^] (int u1'))" by (simp add: int_pow_neg) also have "... = \<^bold>g [^] s \ inv (\<^bold>g [^] u1')" by (simp add: int_pow_int) ultimately show ?thesis by (simp add: inverse_pow_pow pow_generator_mod ) qed lemma P2_inv_g_s_rewrite: assumes "s < order \" shows "\<^bold>g [^] ((r::nat) * \ * u1 + v1 * \) \ inv \<^bold>g [^] (u1 + (order \ - s)) = \<^bold>g [^] (r * \ * u1 + v1 * \) \ \<^bold>g [^] s \ inv \<^bold>g [^] u1" proof- have in_carrier1: "inv \<^bold>g [^] (u1 + (order \ - s)) \ carrier \" by blast have in_carrier2: "inv \<^bold>g [^] u1 \ carrier \" by simp have in_carrier_3: "\<^bold>g [^] (r * \ * u1 + v1 * \) \ carrier \" by simp have "\<^bold>g [^] (r * \ * u1 + v1 * \) \ (inv \<^bold>g [^] (u1 + (order \ - s))) = \<^bold>g [^] (r * \ * u1 + v1 * \) \ (\<^bold>g [^] s \ inv \<^bold>g [^] u1)" using P2_inv_g_rewrite assms by (simp add: inverse_pow_pow) thus ?thesis using cyclic_group_assoc in_carrier1 in_carrier2 by auto qed lemma P2_e0_rewrite: assumes "s < order \ " shows "(\<^bold>g [^] (r * x + xa), \<^bold>g [^] (r * \ * x + xa * \) \ \<^bold>g [^] x) = (\<^bold>g [^] (r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \), \<^bold>g [^] (r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) \ \<^bold>g [^] ((order \ - s + x) mod order \ + s))" proof- have "\<^bold>g [^] (r * x + xa) = \<^bold>g [^] (r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \)" proof- have "[(r * x + xa) = (r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \)] (mod order \)" proof- have "[(r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \) = (r * ((order \ - s + x)) + (r * s + xa))] (mod order \)" by (metis (no_types, lifting) mod_mod_trivial cong_add cong_def mod_mult_right_eq) hence "[(r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \) = r * (order \ - s) + r * x + r * s + xa] (mod order \)" by (simp add: add.assoc distrib_left) hence "[(r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \) = r * x + r * s + r * (order \ - s) + xa] (mod order \)" by (metis add.assoc add.commute) hence "[(r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \) = r * x + r * s + r * order \ - r * s + xa] (mod order \)" proof - have "[(xa + r * s) mod order \ + r * ((x + (order \ - s)) mod order \) = xa + r * (s + x + (order \ - s))] (mod order \)" by (metis (no_types) \[r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ = r * x + r * s + r * (order \ - s) + xa] (mod order \)\ add.commute distrib_left) then show ?thesis by (simp add: assms add.commute distrib_left order.strict_implies_order) qed hence "[(r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \) = r * x + xa] (mod order \)" proof - have "[(xa + r * s) mod order \ + r * ((x + (order \ - s)) mod order \) = xa + (r * x + r * order \)] (mod order \)" by (metis (no_types) \[r * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ = r * x + r * s + r * order \ - r * s + xa] (mod order \)\ add.commute add.left_commute add_diff_cancel_left') then show ?thesis by (metis (no_types) add.commute cong_add_lcancel_nat cong_def distrib_left mod_add_self2 mod_mult_right_eq) qed then show ?thesis using cong_def by metis qed then show ?thesis using finite_group pow_generator_eq_iff_cong by blast qed moreover have "\<^bold>g [^] (r * \ * x + xa * \) \ \<^bold>g [^] x = \<^bold>g [^] (r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) \ \<^bold>g [^] ((order \ - s + x) mod order \ + s)" proof- have "\<^bold>g [^] (r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = \<^bold>g [^] (r * \ * x + xa * \)" proof- have "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * x + xa * \] (mod order \)" proof- have "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = (r * \ * ((order \ - s) + x) + (r * s + xa) * \)] (mod order \)" by (metis (no_types, lifting) cong_add cong_def mod_mult_left_eq mod_mult_right_eq) hence "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * (order \ - s) + r * \ * x + r * s * \ + xa * \] (mod order \)" by (simp add: add.assoc distrib_left distrib_right) hence "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * x + r * s * \ + r * \ * (order \ - s) + xa * \] (mod order \)" by (simp add: add.commute add.left_commute) hence "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * x + r * s * \ + r * \ * order \ - r * \ * s + xa * \] (mod order \)" proof - have "\n na. \ (n::nat) \ na \ n + (na - n) = na" by (meson ordered_cancel_comm_monoid_diff_class.add_diff_inverse) then have "r * \ * s + r * \ * (order \ - s) = r * \ * order \" by (metis add_mult_distrib2 assms less_or_eq_imp_le) then have "r * \ * x + r * s * \ + r * \ * order \ = r * \ * s + r * \ * (order \ - s) + (r * \ * x + r * s * \)" by presburger then have f1: "r * \ * x + r * s * \ + r * \ * order \ - r * \ * s = r * \ * s + r * \ * (order \ - s) - r * \ * s + (r * \ * x + r * s * \)" by simp have "r * \ * s + r * \ * (order \ - s) = r * \ * (order \ - s) + r * \ * s" by presburger then have "r * \ * x + r * s * \ + r * \ * order \ - r * \ * s = r * \ * x + r * s * \ + r * \ * (order \ - s)" using f1 diff_add_inverse2 by presburger then show ?thesis using \[r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \ = r * \ * x + r * s * \ + r * \ * (order \ - s) + xa * \] (mod order \)\ by presburger qed hence "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * x + r * \ * s + r * \ * order \ - r * \ * s + xa * \] (mod order \)" using add.commute add.assoc by force hence "[(r * \ * ((order \ - s + x) mod order \) + (r * s + xa) mod order \ * \) = r * \ * x + r * \ * order \ + xa * \] (mod order \)" by simp thus ?thesis using cong_def semiring_normalization_rules(23) by (simp add: \\c b a. [b = c] (mod a) = (b mod a = c mod a)\ \\c b a. a + b + c = a + c + b\) qed thus ?thesis using finite_group pow_generator_eq_iff_cong by blast qed also have "\<^bold>g [^] ((order \ - s + x) mod order \ + s) = \<^bold>g [^] x" proof- have "[((order \ - s + x) mod order \ + s) = x] (mod order \)" proof- have "[((order \ - s + x) mod order \ + s) = (order \ - s + x+ s)] (mod order \)" by (simp add: add.commute cong_def mod_add_right_eq) hence "[((order \ - s + x) mod order \ + s) = (order \ + x)] (mod order \)" using assms by auto thus ?thesis by (simp add: cong_def) qed thus ?thesis using finite_group pow_generator_eq_iff_cong by blast qed ultimately show ?thesis by argo qed ultimately show ?thesis by simp qed lemma P2_case_l_new_1_gt_e0_rewrite: assumes "s < order \" shows "(\<^bold>g [^] (r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \), \<^bold>g [^] (r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \) \ \<^bold>g [^] (t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \))))) = (\<^bold>g [^] (r * x + xa), \<^bold>g [^] (r * \ * x + xa * \) \ \<^bold>g [^] (t * x))" proof- have "\<^bold>g [^] ((r::nat) * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \) = \<^bold>g [^] (r * x + xa)" proof(cases "r = 0") case True then show ?thesis by (simp add: pow_generator_mod) next case False have "[(r::nat) * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ = r * x + xa] (mod order \)" proof- have "[r * ((order \ * order \ - s * (nat (fst (bezw t (order \))) mod order \) + x) mod order \) + (r * s * nat (fst (bezw t (order \))) + xa) mod order \ = (r * (((order \ * order \ - s * (nat (fst (bezw t (order \))) mod order \)) + x)) + (r * s * nat (fst (bezw t (order \))) + xa))] (mod order \)" proof - have "order \ \ 0" using order_gt_0 by simp then show ?thesis using cong_add cong_def mod_mult_right_eq by (smt mod_mod_trivial) qed hence "[r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ = r * (order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \))) + r * x + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa)] (mod order \)" proof - have "[r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) = r * (order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \))) + r * x] (mod order \)" by (simp add: cong_def distrib_left mod_mult_right_eq) then show ?thesis using assms cong_add gr_implies_not0 by fastforce qed hence "[r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ = r * order \ * order \ - r * s * (nat ((fst (bezw t (order \))) mod order \)) + r * x + r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa] (mod order \)" by (simp add: ab_semigroup_mult_class.mult_ac(1) right_diff_distrib' add.assoc) hence "[r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ = r * order \ * order \ + r * x + xa] (mod order \)" proof- have "r * order \ * order \ - r * s * (nat ((fst (bezw t (order \))) mod order \)) > 0" proof- have "order \ * order \ > s * (nat ((fst (bezw t (order \))) mod order \))" proof- have "(nat ((fst (bezw t (order \))) mod order \)) \ order \" proof - have "\x0 x1. ((x0::int) mod x1 < x1) = (\ x1 + - 1 * (x0 mod x1) \ 0)" by linarith then have "\ int (order \) + - 1 * (fst (bezw t (order \)) mod int (order \)) \ 0" using of_nat_0_less_iff order_gt_0 by fastforce then show ?thesis by linarith qed thus ?thesis using assms proof - have "\n na. \ n \ na \ \ na * order \ < n * nat (fst (bezw t (order \)) mod int (order \))" by (meson \nat (fst (bezw (t::nat) (order \)) mod int (order \)) \ order \\ mult_le_mono not_le) then show ?thesis by(metis (no_types, hide_lams) \(s::nat) < order \\ mult_less_cancel2 nat_less_le not_le not_less_zero) qed qed thus ?thesis using False by auto qed thus ?thesis proof - have "r * order \ * order \ + r * x + xa = r * (order \ * order \ - s * nat (fst (bezw t (order \)) mod int (order \))) + (r * s * nat (fst (bezw t (order \)) mod int (order \)) + xa) + r * x" using \(0::nat) < (r::nat) * order \ * order \ - r * (s::nat) * nat (fst (bezw (t::nat) (order \)) mod int (order \))\ diff_mult_distrib2 by force then show ?thesis by (metis (no_types) \[(r::nat) * ((order \ * order \ - (s::nat) * nat (fst (bezw (t::nat) (order \)) mod int (order \)) + (x::nat)) mod order \) + (r * s * nat (fst (bezw t (order \)) mod int (order \)) + (xa::nat)) mod order \ = r * (order \ * order \ - s * nat (fst (bezw t (order \)) mod int (order \))) + r * x + (r * s * nat (fst (bezw t (order \)) mod int (order \)) + xa)] (mod order \)\ semiring_normalization_rules(23)) qed qed hence "[r * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ = r * x + xa] (mod order \)" by (metis (no_types, lifting) mod_mult_self4 add.assoc mult.commute cong_def) thus ?thesis by blast qed then show ?thesis using finite_group pow_generator_eq_iff_cong by blast qed moreover have "\<^bold>g [^] (r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \) \ \<^bold>g [^] (t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \)))) = \<^bold>g [^] (r * \ * x + xa * \) \ \<^bold>g [^] (t * x)" proof- have "\<^bold>g [^] (r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \) = \<^bold>g [^] (r * \ * x + xa * \)" proof- have "[r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \ = r * \ * x + xa * \] (mod order \)" proof- have "[r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \ = r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \))) + x) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) * \] (mod order \)" proof - show ?thesis by (meson cong_def mod_add_cong mod_mult_left_eq mod_mult_right_eq) qed hence mod_eq: "[r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \ = r * \ * (order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \))) + r * \ * x + r * s * (nat ((fst (bezw t (order \))) mod order \)) * \ + xa * \] (mod order \)" by (simp add: distrib_left distrib_right add.assoc) hence mod_eq': "[r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \ = r * \ * order \ * order \ - r * \ * s * (nat ((fst (bezw t (order \))) mod order \)) + r * \ * x + r * \ * s * (nat ((fst (bezw t (order \))) mod order \)) + xa * \] (mod order \)" by (simp add: semiring_normalization_rules(16) diff_mult_distrib2 semiring_normalization_rules(18)) hence "[r * \ * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \) + (r * s * (nat ((fst (bezw t (order \))) mod order \)) + xa) mod order \ * \ = r * \ * order \ * order \ + r * \ * x + xa * \] (mod order \)" proof(cases "r * \ = 0") case True then show ?thesis by (metis mod_eq' diff_zero mult_0 plus_nat.add_0) next case False hence bound: " r * \ * order \ * order \ - r * \ * s * (nat ((fst (bezw t (order \))) mod order \)) > 0" proof- have "s * (nat ((fst (bezw t (order \))) mod order \)) < order \ * order \" using assms by (simp add: mult_strict_mono nat_less_iff) thus ?thesis using False by auto qed thus ?thesis proof - have "r * \ * order \ * order \ = r * \ * (order \ * order \ - s * nat (fst (bezw t (order \)) mod int (order \))) + r * s * nat (fst (bezw t (order \)) mod int (order \)) * \" using bound diff_mult_distrib2 by force then have "r * \ * order \ * order \ + r * \ * x = r * \ * (order \ * order \ - s * nat (fst (bezw t (order \)) mod int (order \))) + r * \ * x + r * s * nat (fst (bezw t (order \)) mod int (order \)) * \" by presburger then show ?thesis using mod_eq by presburger qed qed thus ?thesis by (metis (mono_tags, lifting) add.assoc cong_def mod_mult_self3) qed then show ?thesis using finite_group pow_generator_eq_iff_cong by blast qed also have "\<^bold>g [^] (t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \)))) = \<^bold>g [^] (t * x)" proof- have "[t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \))) = t * x] (mod order \)" proof- have "[t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \))) = (t * (order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x + s * (nat ((fst (bezw t (order \))) mod order \))))] (mod order \)" using cong_def mod_add_left_eq mod_mult_cong by blast hence "[t * ((order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) + x) mod order \ + s * (nat ((fst (bezw t (order \))) mod order \))) = t * (order \ * order \ + x )] (mod order \)" proof- have "order \ * order \ - s * (nat ((fst (bezw t (order \))) mod order \)) > 0" proof- have "(nat ((fst (bezw t (order \))) mod order \)) \ order \" using nat_le_iff order.strict_implies_order order_gt_0 by (simp add: order.strict_implies_order) thus ?thesis by (metis assms diff_mult_distrib le0 linorder_neqE_nat mult_strict_mono not_le zero_less_diff) qed thus ?thesis using \[(t::nat) * ((order \ * order \ - (s::nat) * nat (fst (bezw t (order \)) mod int (order \)) + (x::nat)) mod order \ + s * nat (fst (bezw t (order \)) mod int (order \))) = t * (order \ * order \ - s * nat (fst (bezw t (order \)) mod int (order \)) + x + s * nat (fst (bezw t (order \)) mod int (order \)))] (mod order \)\ by auto qed thus ?thesis by (metis (no_types, hide_lams) add.commute cong_def mod_mult_right_eq mod_mult_self1) qed thus ?thesis using finite_group pow_generator_eq_iff_cong by blast qed ultimately show ?thesis by argo qed ultimately show ?thesis by simp qed lemma P2_case_l_neq_1_gt_x0_rewrite: assumes "t < order \" and "t \ 0" shows "\<^bold>g [^] (t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \))))) = \<^bold>g [^] (t * u0) \ \<^bold>g [^] s" proof- from assms have gcd: "gcd t (order \) = 1" using prime_field coprime_imp_gcd_eq_1 by blast hence inverse_t: "[ s * (t * (fst (bezw t (order \)))) = s * 1] (mod order \)" by (metis Num.of_nat_simps(2) Num.of_nat_simps(5) cong_scalar_left order_gt_0 inverse) hence inverse_t': "[t * u0 + s * (t * (fst (bezw t (order \)))) =t * u0 + s * 1] (mod order \)" using cong_add_lcancel by fastforce have eq: "int (nat ((fst (bezw t (order \))) mod order \)) = (fst (bezw t (order \))) mod order \" proof- have "(fst (bezw t (order \))) mod order \ \ 0" using order_gt_0 by simp hence "(nat ((fst (bezw t (order \))) mod order \)) = (fst (bezw t (order \))) mod order \" by linarith thus ?thesis by blast qed have "[(t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \))))) = t * u0 + s] (mod order \)" proof- have "[t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \)))) = t * u0 + t * (s * (nat ((fst (bezw t (order \))) mod order \)))] (mod order \)" by (simp add: distrib_left) hence "[t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \)))) = t * u0 + s * (t * (nat ((fst (bezw t (order \))) mod order \)))] (mod order \)" by (simp add: ab_semigroup_mult_class.mult_ac(1) mult.left_commute) hence "[t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \)))) = t * u0 + s * (t * ( ((fst (bezw t (order \))) mod order \)))] (mod order \)" using eq by (simp add: distrib_left mult.commute semiring_normalization_rules(18)) hence "[t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \)))) = t * u0 + s * (t * (fst (bezw t (order \))))] (mod order \)" by (metis (no_types, hide_lams) cong_def mod_add_right_eq mod_mult_right_eq) hence "[t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \)))) = t * u0 + s * 1] (mod order \)" using inverse_t' using cong_trans cong_int_iff by blast thus ?thesis by simp qed hence "\<^bold>g [^] (t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \))))) = \<^bold>g [^] (t * u0 + s)" using finite_group pow_generator_eq_iff_cong by blast thus ?thesis by (simp add: nat_pow_mult) qed text \ Now we show the two end definitions are equal when the input for l (in the ideal model, the second input) is the one constructed by the simulator \ lemma P2_ideal_real_end_eq: assumes b0_inv_b1: "b0 \ inv b1 = (h0 \ inv h1) [^] r" and assert_in_carrier: "h0 \ carrier \ \ h1 \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \" and x1_in_carrier: "x1 \ carrier \" and x0_in_carrier: "x0 \ carrier \" shows "P2_ideal_model_end (x0,x1) (b0 \ (inv (h0 [^] r))) ((h0,h1, \<^bold>g [^] (r::nat),b0,b1),s') \3 = P2_real_model_end (x0,x1) ((h0,h1, \<^bold>g [^] (r::nat),b0,b1),s') \3" including monad_normalisation proof(cases "(b0 \ (inv (h0 [^] r))) = \") \ \ The case distinctions follow the 3 cases give on p 193/194*) \ case True have b1_h1: "b1 = h1 [^] r" proof- from b0_inv_b1 assert_in_carrier have "b0 \ inv b1 = h0 [^] r \ inv h1 [^] r" by (simp add: pow_mult_distrib cyclic_group_commute monoid_comm_monoidI) hence "b0 \ inv h0 [^] r = b1 \ inv h1 [^] r" by (metis Units_eq Units_l_cancel local.inv_equality True assert_in_carrier cyclic_group.inverse_pow_pow cyclic_group_axioms inv_closed nat_pow_closed r_inv) with True have "\ = b1 \ inv h1 [^] r" by (simp add: assert_in_carrier inverse_pow_pow) hence "\ \ h1 [^] r = b1" by (metis assert_in_carrier cyclic_group.inverse_pow_pow cyclic_group_axioms inv_closed inv_inv l_one local.inv_equality nat_pow_closed) thus ?thesis using assert_in_carrier l_one by blast qed obtain \ :: nat where \: "\<^bold>g [^] \ = h1" and "\ < order \" by (metis mod_less_divisor assert_in_carrier generatorE order_gt_0 pow_generator_mod) obtain s :: nat where s: "\<^bold>g [^] s = x1" and s_lt: "s < order \" by (metis assms(3) mod_less_divisor generatorE order_gt_0 pow_generator_mod) have "b1 \ inv \<^bold>g = \<^bold>g [^] (r * \) \ inv \<^bold>g" by (metis \ b1_h1 generator_closed mult.commute nat_pow_pow) have a_g_exp_rewrite: "(\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0 = \<^bold>g [^] (r * u0 + v0)" for u0 v0 by (simp add: nat_pow_mult nat_pow_pow) have z1_rewrite: "(b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ \ = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] u1" for u1 v1 :: nat by (smt \ b1_h1 pow_mult_distrib cyclic_group_commute generator_closed inv_closed m_assoc m_closed monoid_comm_monoidI mult.commute nat_pow_closed nat_pow_mult nat_pow_pow r_one) have z1_rewrite': "\<^bold>g [^] (r * \ * u1 + v1 * \) \ \<^bold>g [^] s \ inv \<^bold>g [^] u1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1" for u1 v1 using assert_in_carrier cyclic_group_commute m_assoc s z1_rewrite by auto have "P2_ideal_model_end (x0,x1) (b0 \ (inv (h0 [^] r))) ((h0,h1, \<^bold>g [^] (r::nat),b0,b1),s') \3 = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ \; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: P2_ideal_model_end_def True funct_OT_12_def) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] u1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: z1_rewrite) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = \<^bold>g [^] (r * u1 + v1); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] u1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: a_g_exp_rewrite) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ map_spmf (\ u1'. (s + u1') mod (order \)) (sample_uniform (order \)); v1 \ map_spmf (\ v1'. ((r * order \ - r * s) + v1') mod (order \)) (sample_uniform (order \)); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = \<^bold>g [^] (r * u1 + v1); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] (u1 + (order \ - s)); let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" apply(simp add: bind_map_spmf o_def Let_def) using P2_output_rewrite assms s_lt assms by presburger also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = \<^bold>g [^] (r * u1 + v1); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = \<^bold>g [^] (r * \ * u1 + v1 * \) \ inv \<^bold>g [^] (u1 + (order \ - s)); let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: samp_uni_plus_one_time_pad) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = \<^bold>g [^] (r * u1 + v1); let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = \<^bold>g [^] (r * \ * u1 + v1 * \) \ \<^bold>g [^] s \ inv \<^bold>g [^] u1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: P2_inv_g_s_rewrite assms s_lt cong: bind_spmf_cong_simp) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: a_g_exp_rewrite z1_rewrite') ultimately show ?thesis by(simp add: P2_real_model_end_def) next obtain \ :: nat where \: "\<^bold>g [^] \ = h0" using generatorE assms using assert_in_carrier by auto have w0_rewrite: "\<^bold>g [^] (r * u0 + v0) = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0" for u0 v0 by (simp add: nat_pow_mult nat_pow_pow) have order_gt_0: "order \ > 0" using order_gt_0 by simp obtain s :: nat where s: "\<^bold>g [^] s = x0" and s_lt: "s < order \" by (metis mod_less_divisor generatorE order_gt_0 pow_generator_mod x0_in_carrier) case False \ \case 2\ hence l_neq_1: "(b0 \ (inv (h0 [^] r))) \ \"by auto then show ?thesis proof(cases "(b0 \ (inv (h0 [^] r))) = \<^bold>g") case True hence "b0 = \<^bold>g \ h0 [^] r" by (metis assert_in_carrier generator_closed inv_solve_right nat_pow_closed) hence "b0 = \<^bold>g \ \<^bold>g [^] (r * \)" by (metis \ generator_closed mult.commute nat_pow_pow) have z0_rewrite: "b0 [^] u0 \ h0 [^] v0 \ \ = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] u0" for u0 v0 :: nat by (smt \ \b0 = \<^bold>g \ \<^bold>g [^] (r * \)\ pow_mult_distrib cyclic_group_commute generator_closed m_assoc monoid_comm_monoidI mult.commute nat_pow_closed nat_pow_mult nat_pow_pow r_one) have z0_rewrite': "\<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (u0 + s) = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] u0 \ \<^bold>g [^] s" for u0 v0 by (simp add: add.assoc nat_pow_mult) have z0_rewrite'': "\<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] u0 \ x0 = b0 [^] u0 \ h0 [^] v0 \ x0" for u0 v0 using z0_rewrite using assert_in_carrier by auto have "P2_ideal_model_end (x0,x1) (b0 \ (inv (h0 [^] r))) ((h0,h1,\<^bold>g [^] (r::nat),b0,b1),s') \3 = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ \; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" apply(simp add: P2_ideal_model_end_def True funct_OT_12_def) using order_gt_0 order_gt_1_gen_not_1 True l_neq_1 by auto also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] u0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: z0_rewrite w0_rewrite) also have "... = do { u0 \ map_spmf (\ u0. ((order \ - s) + u0) mod (order \)) (sample_uniform (order \)); v0 \ map_spmf (\ v0. (r * s + v0) mod (order \)) (sample_uniform (order \)); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (u0 + s); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" apply(simp add: bind_map_spmf o_def Let_def cong: bind_spmf_cong_simp) using P2_e0_rewrite assms s_lt assms by presburger also have "... = do { u0 \ map_spmf (\ u0. ((order \ - s) + u0) mod (order \)) (sample_uniform (order \)); v0 \ map_spmf (\ v0. (r * s + v0) mod (order \)) (sample_uniform (order \)); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] u0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: z0_rewrite' s) also have "... = do { u0 \ map_spmf (\ u0. ((order \ - s) + u0) mod (order \)) (sample_uniform (order \)); v0 \ map_spmf (\ v0. (r * s + v0) mod (order \)) (sample_uniform (order \)); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: w0_rewrite z0_rewrite'') also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: samp_uni_plus_one_time_pad) ultimately show ?thesis by(simp add: P2_real_model_end_def) next case False \ \case 3\ have b0_l: "b0 = (b0 \ (inv (h0 [^] r))) \ h0 [^] r" by (simp add: assert_in_carrier m_assoc) have b0_g_r:"b0 = (b0 \ (inv (h0 [^] r))) \ \<^bold>g [^] (r * \)" by (metis \ b0_l generator_closed mult.commute nat_pow_pow) obtain t :: nat where t: "\<^bold>g [^] t = (b0 \ (inv (h0 [^] r)))" and t_lt_order_g: "t < order \" by (metis (full_types) mod_less_divisor order_gt_0 pow_generator_mod assert_in_carrier cyclic_group.generatorE cyclic_group_axioms inv_closed m_closed nat_pow_closed) with l_neq_1 have t_neq_0: "t \ 0" using l_neq_1_exp_neq_0 by simp have z0_rewrite: "b0 [^] u0 \ h0 [^] v0 \ \ = \<^bold>g [^] (r * \ * u0 + v0 * \) \ ((b0 \ (inv (h0 [^] r)))) [^] u0" for u0 v0 proof- from b0_l have "b0 [^] u0 \ h0 [^] v0 = ((b0 \ (inv (h0 [^] r))) \ h0 [^] r) [^] u0 \ h0 [^] v0" by simp hence "b0 [^] u0 \ h0 [^] v0 = ((b0 \ (inv (h0 [^] r)))) [^] u0 \ (h0 [^] r) [^] u0 \ h0 [^] v0" by (simp add: assert_in_carrier pow_mult_distrib cyclic_group_commute monoid_comm_monoidI) hence "b0 [^] u0 \ h0 [^] v0 = ((\<^bold>g [^] \) [^] r) [^] u0 \ (\<^bold>g [^] \) [^] v0 \ ((b0 \ (inv (h0 [^] r)))) [^] u0" using cyclic_group_assoc cyclic_group_commute assert_in_carrier \ by simp hence "b0 [^] u0 \ h0 [^] v0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ ((b0 \ (inv (h0 [^] r)))) [^] u0" by (simp add: monoid.nat_pow_pow mult.commute nat_pow_mult) thus ?thesis by (simp add: assert_in_carrier) qed have z0_rewrite': "\<^bold>g [^] (r * \ * u0 + v0 * \) \ ((b0 \ (inv (h0 [^] r)))) [^] u0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * u0)" for u0 v0 by (metis generator_closed nat_pow_pow t) have z0_rewrite'': "\<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * u0) \ \<^bold>g [^] s = b0 [^] u0 \ h0 [^] v0 \ x0" for u0 v0 using assert_in_carrier s z0_rewrite z0_rewrite' by auto have "P2_ideal_model_end (x0,x1) (b0 \ (inv (h0 [^] r))) ((h0,h1,\<^bold>g [^] (r::nat),b0,b1),s') \3 = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ ((b0 \ (inv (h0 [^] r)))) [^] u0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: P2_ideal_model_end_def l_neq_1 funct_OT_12_def w0_rewrite z0_rewrite) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * u0); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: z0_rewrite') also have "... = do { u0 \ map_spmf (\ u0. (order \ * order \ - (s * ((nat (((fst (bezw t (order \)))) mod (order \))))) + u0) mod (order \)) (sample_uniform (order \)); v0 \ map_spmf (\ v0. (r * s * (nat ((fst (bezw t (order \))) mod order \)) + v0) mod (order \)) (sample_uniform (order \)); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \))))); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: bind_map_spmf o_def Let_def s_lt P2_case_l_new_1_gt_e0_rewrite cong: bind_spmf_cong_simp) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * (u0 + (s * (nat ((fst (bezw t (order \))) mod order \))))); let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: samp_uni_plus_one_time_pad) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = \<^bold>g [^] (r * u0 + v0); let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = \<^bold>g [^] (r * \ * u0 + v0 * \) \ \<^bold>g [^] (t * u0) \ \<^bold>g [^] s; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: P2_case_l_neq_1_gt_x0_rewrite t_lt_order_g t_neq_0 cyclic_group_assoc) also have "... = do { u0 \ sample_uniform (order \); v0 \ sample_uniform (order \); u1 \ sample_uniform (order \); v1 \ sample_uniform (order \); let w0 = (\<^bold>g [^] (r::nat)) [^] u0 \ \<^bold>g [^] v0; let w1 = (\<^bold>g [^] (r::nat)) [^] u1 \ \<^bold>g [^] v1; let z0 = b0 [^] u0 \ h0 [^] v0 \ x0; let z1 = (b1 \ inv \<^bold>g) [^] u1 \ h1 [^] v1 \ x1; let e0 = (w0,z0); let e1 = (w1,z1); out \ \3 e0 e1 s'; return_spmf ((), out)}" by(simp add: w0_rewrite z0_rewrite'') ultimately show ?thesis by(simp add: P2_real_model_end_def) qed qed lemma P2_ideal_real_eq: assumes x1_in_carrier: "x1 \ carrier \" and x0_in_carrier: "x0 \ carrier \" shows "P2_real_model (x0,x1) \ z \ = P2_ideal_model (x0,x1) \ z \" proof- have "P2_real_model' (x0, x1) \ z \ = P2_ideal_model' (x0, x1) \ z \" proof- have 1:"do { let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; let l = b0 \ (inv (h0 [^] r)); P2_ideal_model_end (x0,x1) l ((h0,h1,a,b0,b1),s') \3} = P2_ideal_model' (x0,x1) \ z \" unfolding P2_ideal_model'_def by simp have "P2_real_model' (x0, x1) \ z \ = do { let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; P2_real_model_end (x0, x1) ((h0,h1,a,b0,b1),s') \3}" by(simp add: P2_real_model'_def) also have "... = do { let (\1, \2, \3) = \; ((h0,h1,a,b0,b1),s) \ \1 \ z; _ :: unit \ assert_spmf (h0 \ carrier \ \ h1 \ carrier \ \ a \ carrier \ \ b0 \ carrier \ \ b1 \ carrier \); (((in1, in2, in3), r),s') \ \2 (h0,h1,a,b0,b1) s; let (h,a,b) = (h0 \ inv h1, a, b0 \ inv b1); (out_zk_funct, _) \ funct_DH_ZK (h,a,b) ((in1, in2, in3), r); _ :: unit \ assert_spmf out_zk_funct; let l = b0 \ (inv (h0 [^] r)); P2_ideal_model_end (x0,x1) l ((h0,h1,a,b0,b1),s') \3}" by(simp add: P2_ideal_real_end_eq assms cong: bind_spmf_cong_simp) ultimately show ?thesis by(simp add: P2_real_model'_def P2_ideal_model'_def) qed thus ?thesis by(simp add: P2_ideal_model_rewrite P2_real_model_rewrite) qed lemma malicious_sec_P2: assumes x1_in_carrier: "x1 \ carrier \" and x0_in_carrier: "x0 \ carrier \" - shows "mal_def.inf_theoretic_mal_sec_P2 (x0,x1) \ z (P2_S1, P2_S2) \" - unfolding malicious_base.inf_theoretic_mal_sec_P2_def + shows "mal_def.perfect_sec_P2 (x0,x1) \ z (P2_S1, P2_S2) \" + unfolding malicious_base.perfect_sec_P2_def by (simp add: P2_ideal_real_eq P2_ideal_view_unfold assms) (*correctness*) lemma correct: assumes "x0 \ carrier \" and "x1 \ carrier \" shows "funct_OT_12 (x0, x1) \ = protocol_ot (x0,x1) \" proof- have \_eq_0_output_correct: "((\<^bold>g [^] \0) [^] r) [^] u0 \ (\<^bold>g [^] \0) [^] v0 \ x0 \ inv (((\<^bold>g [^] r) [^] u0 \ \<^bold>g [^] v0) [^] \0) = x0" (is "?lhs = ?rhs") for \0 r u0 v0 :: nat proof- have mult_com: "r * u0 * \0 = \0 * r * u0" by simp have in_carrier1: "((\<^bold>g [^] (r * u0 * \0))) \ (\<^bold>g [^] (v0 * \0)) \ carrier \" by simp have in_carrier2: "inv ((\<^bold>g [^] (r * u0 * \0))) \ (\<^bold>g [^] (v0 * \0)) \ carrier \" by simp have "?lhs = ((\<^bold>g [^] (\0 * r * u0))) \ (\<^bold>g [^] (\0 * v0)) \ x0 \ inv (((\<^bold>g [^] (r * u0 * \0)) \ \<^bold>g [^] (v0 * \0)))" by (simp add: nat_pow_pow pow_mult_distrib cyclic_group_commute monoid_comm_monoidI) also have "... = (((\<^bold>g [^] (r * u0 * \0))) \ (\<^bold>g [^] (v0 * \0))) \ x0 \ (inv (((\<^bold>g [^] (r * u0 * \0)) \ \<^bold>g [^] (v0 * \0))))" using mult.commute mult.assoc mult_com by (metis (no_types) mult.commute) also have "... = x0 \ (((\<^bold>g [^] (r * u0 * \0))) \ (\<^bold>g [^] (v0 * \0))) \ (inv (((\<^bold>g [^] (r * u0 * \0)) \ \<^bold>g [^] (v0 * \0))))" using cyclic_group_commute in_carrier1 assms by simp also have "... = x0 \ ((((\<^bold>g [^] (r * u0 * \0))) \ (\<^bold>g [^] (v0 * \0))) \ (inv (((\<^bold>g [^] (r * u0 * \0)) \ \<^bold>g [^] (v0 * \0)))))" using cyclic_group_assoc in_carrier1 in_carrier2 assms by auto ultimately show ?thesis using assms by simp qed have \_eq_1_output_correct: "((\<^bold>g [^] \1) [^] r \ \<^bold>g \ inv \<^bold>g) [^] u1 \ (\<^bold>g [^] \1) [^] v1 \ x1 \ inv (((\<^bold>g [^] r) [^] u1 \ \<^bold>g [^] v1) [^] \1) = x1" (is "?lhs = ?rhs") for \1 r u1 v1 :: nat proof- have com1: "\1 * r * u1 = r * u1 * \1" "v1 * \1 = \1 * v1" by simp+ have in_carrier1: "(\<^bold>g [^] (r * u1 * \1)) \ (\<^bold>g [^] (v1 * \1)) \ carrier \" by simp have in_carrier2: "inv ((\<^bold>g [^] (r * u1 * \1)) \ (\<^bold>g [^] (v1 * \1))) \ carrier \" by simp have lhs: "?lhs = ((\<^bold>g [^] (\1*r)) \ \<^bold>g \ inv \<^bold>g) [^] u1 \ (\<^bold>g [^] (\1 * v1)) \ x1 \ inv ((\<^bold>g [^] (r * u1 * \1)) \ \<^bold>g [^] (v1*\1))" by (simp add: nat_pow_pow pow_mult_distrib cyclic_group_commute monoid_comm_monoidI) also have lhs1: "... = (\<^bold>g [^] (\1 * r)) [^] u1 \ (\<^bold>g [^] (\1 * v1)) \ x1 \ inv ((\<^bold>g [^] (r * u1 * \1)) \ \<^bold>g [^] (v1*\1))" by (simp add: cyclic_group_assoc) also have lhs2: "... = (\<^bold>g [^] (r * u1 * \1)) \ (\<^bold>g [^] (v1 * \1)) \ x1 \ inv ((\<^bold>g [^] (r * u1 * \1)) \ \<^bold>g [^] (v1 * \1))" by (simp add: nat_pow_pow pow_mult_distrib cyclic_group_commute monoid_comm_monoidI com1) also have "... = (((\<^bold>g [^] (r * u1 * \1)) \ (\<^bold>g [^] (v1 * \1))) \ x1) \ inv ((\<^bold>g [^] (r * u1 * \1)) \ \<^bold>g [^] (v1 * \1))" using in_carrier1 in_carrier2 assms cyclic_group_assoc by blast also have "... = (x1 \ ((\<^bold>g [^] (r * u1 * \1)) \ (\<^bold>g [^] (v1 * \1)))) \ inv ((\<^bold>g [^] (r * u1 * \1)) \ \<^bold>g [^] (v1 * \1))" using in_carrier1 assms cyclic_group_commute by simp ultimately show ?thesis using cyclic_group_assoc assms in_carrier1 in_carrier1 assms cyclic_group_commute lhs1 lhs2 lhs by force qed show ?thesis unfolding funct_OT_12_def protocol_ot_def Let_def by(cases \; auto simp add: assms \_eq_1_output_correct \_eq_0_output_correct bind_spmf_const lossless_sample_uniform_units order_gt_0 P1_assert_correct1 P1_assert_correct2 lossless_weight_spmfD) qed lemma correctness: assumes "x0 \ carrier \" and "x1 \ carrier \" shows "mal_def.correct (x0,x1) \" unfolding mal_def.correct_def by(simp add: correct assms) end locale OT_asymp = fixes \ :: "nat \ 'grp cyclic_group" assumes ot: "\\. ot (\ \)" begin sublocale ot "\ n" for n using ot by simp lemma correctness_asym: assumes "x0 \ carrier (\ n)" and "x1 \ carrier (\ n)" shows "mal_def.correct n (x0,x1) \" using assms correctness by simp lemma P1_security_asym: "negligible (\ n. mal_def.adv_P1 n M \ z (P1_S1 n, P1_S2) \ D)" if neg1: "negligible (\ n. ddh.advantage n (P1_DDH_mal_adv_\_true n M z \ D))" and neg2: "negligible (\ n. ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_true n M z \ D)))" and neg3: "negligible (\ n. ddh.advantage n (P1_DDH_mal_adv_\_false n M z \ D))" and neg4: "negligible (\ n. ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_false n M z \ D)))" proof- have neg_add1: "negligible (\ n. ddh.advantage n (P1_DDH_mal_adv_\_true n M z \ D) + ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_true n M z \ D)))" and neg_add2: "negligible (\ n. ddh.advantage n (P1_DDH_mal_adv_\_false n M z \ D) + ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_false n M z \ D)))" using neg1 neg2 neg3 neg4 negligible_plus by(blast)+ show ?thesis proof(cases \) case True have bound_mod: "\mal_def.adv_P1 n M \ z (P1_S1 n, P1_S2) \ D\ \ ddh.advantage n (P1_DDH_mal_adv_\_true n M z \ D) + ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_true n M z \ D))" for n by (metis (no_types) True abs_idempotent P1_adv_real_ideal_model_def P1_advantages_eq P1_real_ideal_DDH_advantage_true_bound) then show ?thesis using P1_real_ideal_DDH_advantage_true_bound that bound_mod that negligible_le neg_add1 by presburger next case False have bound_mod: "\mal_def.adv_P1 n M \ z (P1_S1 n, P1_S2) \ D\ \ ddh.advantage n (P1_DDH_mal_adv_\_false n M z \ D) + ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_false n M z \ D))" for n proof - have "\spmf (P1_real_model n M \ z \ \ D) True - spmf (P1_ideal_model n M \ z \ \ D) True\ \ local.ddh.advantage n (P1_DDH_mal_adv_\_false n M z \ D) + local.ddh.advantage n (ddh.DDH_\' n (P1_DDH_mal_adv_\_false n M z \ D))" by (metis (no_types) False P1_adv_real_ideal_model_def P1_advantages_eq P1_real_ideal_DDH_advantage_false_bound) then show ?thesis by (simp add: P1_adv_real_ideal_model_def P1_advantages_eq) qed then show ?thesis using P1_real_ideal_DDH_advantage_false_bound bound_mod that negligible_le neg_add2 by presburger qed qed lemma P2_security_asym: assumes x1_in_carrier: "x1 \ carrier (\ n)" and x0_in_carrier: "x0 \ carrier (\ n)" - shows "mal_def.inf_theoretic_mal_sec_P2 n (x0,x1) \ z (P2_S1 n, P2_S2 n) \" + shows "mal_def.perfect_sec_P2 n (x0,x1) \ z (P2_S1 n, P2_S2 n) \" using assms malicious_sec_P2 by fast end end diff --git a/thys/Multi_Party_Computation/Noar_Pinkas_OT.thy b/thys/Multi_Party_Computation/Noar_Pinkas_OT.thy --- a/thys/Multi_Party_Computation/Noar_Pinkas_OT.thy +++ b/thys/Multi_Party_Computation/Noar_Pinkas_OT.thy @@ -1,363 +1,363 @@ subsection \Noar Pinkas OT\ text\Here we prove security for the Noar Pinkas OT from \cite{DBLP:conf/soda/NaorP01}.\ theory Noar_Pinkas_OT imports Cyclic_Group_Ext Game_Based_Crypto.Diffie_Hellman OT_Functionalities Semi_Honest_Def Uniform_Sampling begin locale np_base = fixes \ :: "'grp cyclic_group" (structure) assumes finite_group: "finite (carrier \)" and or_gt_0: "0 < order \" and prime_order: "prime (order \)" begin lemma prime_field: "a < (order \) \ a \ 0 \ coprime a (order \)" by(metis dvd_imp_le neq0_conv not_le prime_imp_coprime prime_order coprime_commute) lemma weight_sample_uniform_units: "weight_spmf (sample_uniform_units (order \)) = 1" using lossless_spmf_def lossless_sample_uniform_units prime_order prime_gt_1_nat by auto definition protocol :: "('grp \ 'grp) \ bool \ (unit \ 'grp) spmf" where "protocol M v = do { let (m0,m1) = M; a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); c\<^sub>v' :: nat \ sample_uniform (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; let z0' = ((\<^bold>g [^] (if v then c\<^sub>v' else c\<^sub>v)) [^] s0) \ ((\<^bold>g [^] b) [^] r0); r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z1' = ((\<^bold>g [^] ((if v then c\<^sub>v else c\<^sub>v'))) [^] s1) \ ((\<^bold>g [^] b) [^] r1); let enc_m0 = z0' \ m0; let enc_m1 = z1' \ m1; let out_2 = (if v then enc_m1 \ inv (w1 [^] b) else enc_m0 \ inv (w0 [^] b)); return_spmf ((), out_2)}" lemma lossless_protocol: "lossless_spmf (protocol M \)" apply(auto simp add: protocol_def Let_def split_def lossless_sample_uniform_units or_gt_0) using prime_order prime_gt_1_nat lossless_sample_uniform_units by simp type_synonym 'grp' view1 = "(('grp' \ 'grp') \ ('grp' \ 'grp' \ 'grp' \ 'grp')) spmf" type_synonym 'grp' dist_adversary = "(('grp' \ 'grp') \ 'grp' \ 'grp' \ 'grp' \ 'grp') \ bool spmf" definition R1 :: "('grp \ 'grp) \ bool \ 'grp view1" where "R1 msgs \ = do { let (m0, m1) = msgs; a \ sample_uniform (order \); b \ sample_uniform (order \); let c\<^sub>\ = a*b; c\<^sub>\' \ sample_uniform (order \); return_spmf (msgs, (\<^bold>g [^] a, \<^bold>g [^] b, (if \ then \<^bold>g [^] c\<^sub>\' else \<^bold>g [^] c\<^sub>\), (if \ then \<^bold>g [^] c\<^sub>\ else \<^bold>g [^] c\<^sub>\')))}" lemma lossless_R1: "lossless_spmf (R1 M \)" by(simp add: R1_def Let_def lossless_sample_uniform_units or_gt_0) definition inter :: "('grp \ 'grp) \ 'grp view1" where "inter msgs = do { a \ sample_uniform (order \); b \ sample_uniform (order \); c \ sample_uniform (order \); d \ sample_uniform (order \); return_spmf (msgs, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c, \<^bold>g [^] d)}" definition S1 :: "('grp \ 'grp) \ unit \ 'grp view1" where "S1 msgs out1 = do { let (m0, m1) = msgs; a \ sample_uniform (order \); b \ sample_uniform (order \); c \ sample_uniform (order \); return_spmf (msgs, (\<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c, \<^bold>g [^] (a*b)))}" lemma lossless_S1: "lossless_spmf (S1 M out1)" by(simp add: S1_def Let_def lossless_sample_uniform_units or_gt_0) fun R1_inter_adversary :: "'grp dist_adversary \ ('grp \ 'grp) \ 'grp \ 'grp \ 'grp \ bool spmf" where "R1_inter_adversary \ msgs \ \ \ = do { c \ sample_uniform (order \); \ (msgs, \, \, \, \<^bold>g [^] c)}" fun inter_S1_adversary :: "'grp dist_adversary \ ('grp \ 'grp) \ 'grp \ 'grp \ 'grp \ bool spmf" where "inter_S1_adversary \ msgs \ \ \ = do { c \ sample_uniform (order \); \ (msgs, \, \, \<^bold>g [^] c, \)}" sublocale ddh: ddh \ . definition R2 :: "('grp \ 'grp) \ bool \ (bool \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp) spmf" where "R2 M v = do { let (m0,m1) = M; a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); c\<^sub>v' :: nat \ sample_uniform (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; let z = ((\<^bold>g [^] c\<^sub>v') [^] s0) \ ((\<^bold>g [^] b) [^] r0); r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); let enc_m = z \ (if v then m0 else m1); let enc_m' = z' \ (if v then m1 else m0) ; return_spmf(v, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" lemma lossless_R2: "lossless_spmf (R2 M \)" apply(simp add: R2_def Let_def split_def lossless_sample_uniform_units or_gt_0) using prime_order prime_gt_1_nat lossless_sample_uniform_units by simp definition S2 :: "bool \ 'grp \ (bool \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp \ 'grp) spmf" where "S2 v m = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); s' \ sample_uniform (order \); let enc_m = \<^bold>g [^] s'; let enc_m' = z' \ m ; return_spmf(v, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" lemma lossless_S2: "lossless_spmf (S2 \ out2)" apply(simp add: S2_def Let_def lossless_sample_uniform_units or_gt_0) using prime_order prime_gt_1_nat lossless_sample_uniform_units by simp sublocale sim_def: sim_det_def R1 S1 R2 S2 funct_OT_12 protocol unfolding sim_det_def_def by(auto simp add: lossless_R1 lossless_S1 lossless_R2 lossless_S2 lossless_protocol lossless_funct_OT_12) end locale np = np_base + cyclic_group \ begin lemma protocol_inverse: assumes "m0 \ carrier \" "m1 \ carrier \" shows" ((\<^bold>g [^] ((a*b) mod (order \))) [^] (s1 :: nat)) \ ((\<^bold>g [^] b) [^] (r1::nat)) \ (if v then m0 else m1) \ inv (((\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1) [^] b) = (if v then m0 else m1)" (is "?lhs = ?rhs") proof- have 1: "(a*b)*(s1) + b*r1 =((a::nat)*(s1) + r1)*b " using mult.commute mult.assoc add_mult_distrib by auto have "?lhs = ((\<^bold>g [^] (a*b)) [^] s1) \ ((\<^bold>g [^] b) [^] r1) \ (if v then m0 else m1) \ inv (((\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1) [^] b)" by(simp add: pow_generator_mod) also have "... = (\<^bold>g [^] ((a*b)*(s1))) \ ((\<^bold>g [^] (b*r1))) \ ((if v then m0 else m1) \ inv (((\<^bold>g [^] ((a*(s1) + r1)*b)))))" by(auto simp add: nat_pow_pow nat_pow_mult assms cyclic_group_assoc) also have "... = \<^bold>g [^] ((a*b)*(s1)) \ \<^bold>g [^] (b*r1) \ ((inv (((\<^bold>g [^] ((a*(s1) + r1)*b))))) \ (if v then m0 else m1))" by(simp add: nat_pow_mult cyclic_group_commute assms) also have "... = (\<^bold>g [^] ((a*b)*(s1) + b*r1) \ inv (((\<^bold>g [^] ((a*(s1) + r1)*b))))) \ (if v then m0 else m1)" by(simp add: nat_pow_mult cyclic_group_assoc assms) also have "... = (\<^bold>g [^] ((a*b)*(s1) + b*r1) \ inv (((\<^bold>g [^] (((a*b)*(s1) + r1*b)))))) \ (if v then m0 else m1)" using 1 by (simp add: mult.commute) ultimately show ?thesis using l_cancel_inv assms by (simp add: mult.commute) qed lemma correctness: assumes "m0 \ carrier \" "m1 \ carrier \" shows "sim_def.correctness (m0,m1) \" proof- have "protocol (m0, m1) \ = funct_OT_12 (m0, m1) \" proof- have "protocol (m0, m1) \ = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let out_2 = ((\<^bold>g [^] ((a*b) mod (order \))) [^] s1) \ ((\<^bold>g [^] b) [^] r1) \ (if \ then m1 else m0) \ inv (((\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1) [^] b); return_spmf ((), out_2)}" by(simp add: protocol_def lossless_sample_uniform_units bind_spmf_const weight_sample_uniform_units or_gt_0) also have "... = do { let out_2 = (if \ then m1 else m0); return_spmf ((), out_2)}" by(simp add: protocol_inverse assms lossless_sample_uniform_units bind_spmf_const weight_sample_uniform_units or_gt_0) ultimately show ?thesis by(simp add: Let_def funct_OT_12_def) qed thus ?thesis by(simp add: sim_def.correctness_def) qed lemma security_P1: shows "sim_def.adv_P1 msgs \ D \ ddh.advantage (R1_inter_adversary D msgs) + ddh.advantage (inter_S1_adversary D msgs)" (is "?lhs \ ?rhs") proof(cases \) case True have "R1 msgs \ = S1 msgs out1" for out1 by(simp add: R1_def S1_def True) then have "sim_def.adv_P1 msgs \ D = 0" by(simp add: sim_def.adv_P1_def funct_OT_12_def) also have "ddh.advantage A \ 0" for A using ddh.advantage_def by simp ultimately show ?thesis by simp next case False have bounded_advantage: "\(a :: real) - b\ = e1 \ \b - c\ = e2 \ \a - c\ \ e1 + e2" for a b e1 c e2 by simp also have R1_inter_dist: "\spmf (R1 msgs False \ D) True - spmf ((inter msgs) \ D) True\ = ddh.advantage (R1_inter_adversary D msgs)" unfolding R1_def inter_def ddh.advantage_def ddh.ddh_0_def ddh.ddh_1_def Let_def split_def by(simp) also have inter_S1_dist: "\spmf ((inter msgs) \ D) True - spmf (S1 msgs out1 \ D) True\ = ddh.advantage (inter_S1_adversary D msgs)" for out1 including monad_normalisation by(simp add: S1_def inter_def ddh.advantage_def ddh.ddh_0_def ddh.ddh_1_def) ultimately have "\spmf (R1 msgs False \ (\view. D view)) True - spmf (S1 msgs out1 \ (\view. D view)) True\ \ ?rhs" for out1 using R1_inter_dist by auto thus ?thesis by(simp add: sim_def.adv_P1_def funct_OT_12_def False) qed lemma add_mult_one_time_pad: assumes "s0 < order \" and "s0 \ 0" shows "map_spmf (\ c\<^sub>v'. (((b* r0) + (s0 * c\<^sub>v')) mod(order \))) (sample_uniform (order \)) = sample_uniform (order \)" proof- have "gcd s0 (order \) = 1" using assms prime_field by simp thus ?thesis using add_mult_one_time_pad by force qed lemma security_P2: assumes "m0 \ carrier \" "m1 \ carrier \" - shows "sim_def.inf_theoretic_P2 (m0,m1) \" + shows "sim_def.perfect_sec_P2 (m0,m1) \" proof- have "R2 (m0, m1) \ = S2 \ (if \ then m1 else m0)" including monad_normalisation proof- have "R2 (m0, m1) \ = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); c\<^sub>v' :: nat \ sample_uniform (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; let s' = (((b* r0) + ((c\<^sub>v')*(s0))) mod(order \)); let z = \<^bold>g [^] s' ; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); let enc_m = z \ (if \ then m0 else m1); let enc_m' = z' \ (if \ then m1 else m0) ; return_spmf(\, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" by(simp add: R2_def nat_pow_pow nat_pow_mult pow_generator_mod add.commute) also have "... = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; s' \ map_spmf (\ c\<^sub>v'. (((b* r0) + ((c\<^sub>v')*(s0))) mod(order \))) (sample_uniform (order \)); let z = \<^bold>g [^] s'; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); let enc_m = z \ (if \ then m0 else m1); let enc_m' = z' \ (if \ then m1 else m0) ; return_spmf(\, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; s' \ (sample_uniform (order \)); let z = \<^bold>g [^] s'; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); let enc_m = z \ (if \ then m0 else m1); let enc_m' = z' \ (if \ then m1 else m0) ; return_spmf(\, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" by(simp add: add_mult_one_time_pad Let_def mult.commute cong: bind_spmf_cong_simp) also have "... = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); enc_m \ map_spmf (\ s'. \<^bold>g [^] s' \ (if \ then m0 else m1)) (sample_uniform (order \)); let enc_m' = z' \ (if \ then m1 else m0) ; return_spmf(\, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { a :: nat \ sample_uniform (order \); b :: nat \ sample_uniform (order \); let c\<^sub>v = (a*b) mod (order \); r0 :: nat \ sample_uniform_units (order \); s0 :: nat \ sample_uniform_units (order \); let w0 = (\<^bold>g [^] a) [^] s0 \ \<^bold>g [^] r0; r1 :: nat \ sample_uniform_units (order \); s1 :: nat \ sample_uniform_units (order \); let w1 = (\<^bold>g [^] a) [^] s1 \ \<^bold>g [^] r1; let z' = ((\<^bold>g [^] (c\<^sub>v)) [^] s1) \ ((\<^bold>g [^] b) [^] r1); enc_m \ map_spmf (\ s'. \<^bold>g [^] s') (sample_uniform (order \)); let enc_m' = z' \ (if \ then m1 else m0) ; return_spmf(\, \<^bold>g [^] a, \<^bold>g [^] b, \<^bold>g [^] c\<^sub>v, w0, enc_m, w1, enc_m')}" by(simp add: sample_uniform_one_time_pad assms) ultimately show ?thesis by(simp add: S2_def Let_def bind_map_spmf o_def) qed thus ?thesis - by(simp add: sim_def.inf_theoretic_P2_def funct_OT_12_def) + by(simp add: sim_def.perfect_sec_P2_def funct_OT_12_def) qed end locale np_asymp = fixes \ :: "security \ 'grp cyclic_group" assumes np: "\\. np (\ \)" begin sublocale np "\ \" for \ by(simp add: np) theorem correctness_asymp: assumes "m0 \ carrier (\ \)" "m1 \ carrier (\ \)" shows "sim_def.correctness \ (m0, m1) \" by(simp add: correctness assms) theorem security_P1_asymp: assumes "negligible (\ \. ddh.advantage \ (inter_S1_adversary \ D msgs))" and "negligible (\ \. ddh.advantage \ (R1_inter_adversary \ D msgs))" shows "negligible (\ \. sim_def.adv_P1 \ msgs \ D)" proof- have "sim_def.adv_P1 \ msgs \ D \ ddh.advantage \ (R1_inter_adversary \ D msgs) + ddh.advantage \ (inter_S1_adversary \ D msgs)" for \ using security_P1 by simp moreover have "negligible (\ \. ddh.advantage \ (R1_inter_adversary \ D msgs) + ddh.advantage \ (inter_S1_adversary \ D msgs))" using assms by (simp add: negligible_plus) ultimately show ?thesis using negligible_le sim_def.adv_P1_def by auto qed theorem security_P2_asymp: assumes "m0 \ carrier (\ \)" "m1 \ carrier (\ \)" - shows "sim_def.inf_theoretic_P2 \ (m0,m1) \" + shows "sim_def.perfect_sec_P2 \ (m0,m1) \" by(simp add: security_P2 assms) end end \ No newline at end of file diff --git a/thys/Multi_Party_Computation/OT14.thy b/thys/Multi_Party_Computation/OT14.thy --- a/thys/Multi_Party_Computation/OT14.thy +++ b/thys/Multi_Party_Computation/OT14.thy @@ -1,742 +1,742 @@ subsection \1-out-of-2 OT to 1-out-of-4 OT\ text \Here we construct a protocol that achieves 1-out-of-4 OT from 1-out-of-2 OT. We follow the protocol for constructing 1-out-of-n OT from 1-out-of-2 OT from \cite{DBLP:books/cu/Goldreich2004}. We assume the security properties on 1-out-of-2 OT.\ theory OT14 imports Semi_Honest_Def OT_Functionalities Uniform_Sampling begin type_synonym input1 = "bool \ bool \ bool \ bool" type_synonym input2 = "bool \ bool" type_synonym 'v_OT121' view1 = "(input1 \ (bool \ bool \ bool \ bool \ bool \ bool) \ 'v_OT121' \ 'v_OT121' \ 'v_OT121')" type_synonym 'v_OT122' view2 = "(input2 \ (bool \ bool \ bool \ bool) \ 'v_OT122' \ 'v_OT122' \ 'v_OT122')" locale ot14_base = fixes S1_OT12 :: "(bool \ bool) \ unit \ 'v_OT121 spmf" \ \simulator for party 1 in OT12\ and R1_OT12 :: "(bool \ bool) \ bool \ 'v_OT121 spmf" \ \real view for party 1 in OT12\ and adv_OT12 :: real and S2_OT12 :: "bool \ bool \ 'v_OT122 spmf" and R2_OT12 :: "(bool \ bool) \ bool \ 'v_OT122 spmf" and protocol_OT12 :: "(bool \ bool) \ bool \ (unit \ bool) spmf" assumes ass_adv_OT12: "sim_det_def.adv_P1 R1_OT12 S1_OT12 funct_OT12 (m0,m1) c D \ adv_OT12" \ \bound the advantage of OT12 for party 1\ - and inf_th_OT12_P2: "sim_det_def.inf_theoretic_P2 R2_OT12 S2_OT12 funct_OT12 (m0,m1) \" \ \information theoretic security for party 2\ + and inf_th_OT12_P2: "sim_det_def.perfect_sec_P2 R2_OT12 S2_OT12 funct_OT12 (m0,m1) \" \ \information theoretic security for party 2\ and correct: "protocol_OT12 msgs b = funct_OT_12 msgs b" and lossless_R1_12: "lossless_spmf (R1_OT12 m c)" and lossless_S1_12: "lossless_spmf (S1_OT12 m out1)" and lossless_S2_12: "lossless_spmf (S2_OT12 c out2)" and lossless_R2_12: "lossless_spmf (R2_OT12 M c)" and lossless_funct_OT12: "lossless_spmf (funct_OT12 (m0,m1) c)" and lossless_protocol_OT12: "lossless_spmf (protocol_OT12 M C)" begin sublocale OT_12_sim: sim_det_def R1_OT12 S1_OT12 R2_OT12 S2_OT12 funct_OT_12 protocol_OT12 unfolding sim_det_def_def by(simp add: lossless_R1_12 lossless_S1_12 lossless_funct_OT12 lossless_R2_12 lossless_S2_12) lemma OT_12_P1_assms_bound': "\spmf (bind_spmf (R1_OT12 (m0,m1) c) (\ view. ((D::'v_OT121 \ bool spmf) view ))) True - spmf (bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (D view ))) True\ \ adv_OT12" proof- have "sim_det_def.adv_P1 R1_OT12 S1_OT12 funct_OT_12 (m0,m1) c D = \spmf (bind_spmf (R1_OT12 (m0,m1) c) (\ view. (D view ))) True - spmf (funct_OT_12 (m0,m1) c \ (\ ((out1::unit), (out2::bool)). S1_OT12 (m0,m1) out1 \ (\ view. D view))) True\" using sim_det_def.adv_P1_def using OT_12_sim.adv_P1_def by auto also have "... = \spmf (bind_spmf (R1_OT12 (m0,m1) c) (\ view. ((D::'v_OT121 \ bool spmf) view ))) True - spmf (bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (D view ))) True\" by(simp add: funct_OT_12_def) ultimately show ?thesis by(metis ass_adv_OT12) qed lemma OT_12_P2_assm: "R2_OT12 (m0,m1) \ = funct_OT_12 (m0,m1) \ \ (\ (out1, out2). S2_OT12 \ out2)" - using inf_th_OT12_P2 OT_12_sim.inf_theoretic_P2_def by blast + using inf_th_OT12_P2 OT_12_sim.perfect_sec_P2_def by blast definition protocol_14_OT :: "input1 \ input2 \ (unit \ bool) spmf" where "protocol_14_OT M C = do { let (c0,c1) = C; let (m00, m01, m10, m11) = M; S0 \ coin_spmf; S1 \ coin_spmf; S2 \ coin_spmf; S3 \ coin_spmf; S4 \ coin_spmf; S5 \ coin_spmf; let a0 = S0 \ S2 \ m00; let a1 = S0 \ S3 \ m01; let a2 = S1 \ S4 \ m10; let a3 = S1 \ S5 \ m11; (_,Si) \ protocol_OT12 (S0, S1) c0; (_,Sj) \ protocol_OT12 (S2, S3) c1; (_,Sk) \ protocol_OT12 (S4, S5) c1; let s2 = Si \ (if c0 then Sk else Sj) \ (if c0 then (if c1 then a3 else a2) else (if c1 then a1 else a0)); return_spmf ((), s2)}" lemma lossless_protocol_14_OT: "lossless_spmf (protocol_14_OT M C)" by(simp add: protocol_14_OT_def lossless_protocol_OT12 split_def) definition R1_14 :: "input1 \ input2 \ 'v_OT121 view1 spmf" where "R1_14 msgs choice = do { let (m00, m01, m10, m11) = msgs; let (c0, c1) = choice; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ R1_OT12 (S0, S1) c0; b :: 'v_OT121 \ R1_OT12 (S2, S3) c1; c :: 'v_OT121 \ R1_OT12 (S4, S5) c1; return_spmf (msgs, (S0, S1, S2, S3, S4, S5), a, b, c)}" lemma lossless_R1_14: "lossless_spmf (R1_14 msgs C)" by(simp add: R1_14_def split_def lossless_R1_12) definition R1_14_interm1 :: "input1 \ input2 \ 'v_OT121 view1 spmf" where "R1_14_interm1 msgs choice = do { let (m00, m01, m10, m11) = msgs; let (c0, c1) = choice; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ S1_OT12 (S0, S1) (); b :: 'v_OT121 \ R1_OT12 (S2, S3) c1; c :: 'v_OT121 \ R1_OT12 (S4, S5) c1; return_spmf (msgs, (S0, S1, S2, S3, S4, S5), a, b, c)}" lemma lossless_R1_14_interm1: "lossless_spmf (R1_14_interm1 msgs C)" by(simp add: R1_14_interm1_def split_def lossless_R1_12 lossless_S1_12) definition R1_14_interm2 :: "input1 \ input2 \ 'v_OT121 view1 spmf" where "R1_14_interm2 msgs choice = do { let (m00, m01, m10, m11) = msgs; let (c0, c1) = choice; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ S1_OT12 (S0, S1) (); b :: 'v_OT121 \ S1_OT12 (S2, S3) (); c :: 'v_OT121 \ R1_OT12 (S4, S5) c1; return_spmf (msgs, (S0, S1, S2, S3, S4, S5), a, b, c)}" lemma lossless_R1_14_interm2: "lossless_spmf (R1_14_interm2 msgs C)" by(simp add: R1_14_interm2_def split_def lossless_R1_12 lossless_S1_12) definition S1_14 :: "input1 \ unit \ 'v_OT121 view1 spmf" where "S1_14 msgs _ = do { let (m00, m01, m10, m11) = msgs; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ S1_OT12 (S0, S1) (); b :: 'v_OT121 \ S1_OT12 (S2, S3) (); c :: 'v_OT121 \ S1_OT12 (S4, S5) (); return_spmf (msgs, (S0, S1, S2, S3, S4, S5), a, b, c)}" lemma lossless_S1_14: "lossless_spmf (S1_14 m out)" by(simp add: S1_14_def lossless_S1_12) lemma reduction_step1: shows "\ A1. \spmf (bind_spmf (R1_14 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\" including monad_normalisation proof- define A1' where "A1' == \ (view :: 'v_OT121) (m0,m1). do { S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; b :: 'v_OT121 \ R1_OT12 (S2, S3) c1; c :: 'v_OT121 \ R1_OT12 (S4, S5) c1; let R = (M, (m0,m1, S2, S3, S4, S5), view, b, c); D R}" have "\spmf (bind_spmf (R1_14 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1' view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1' view (m0,m1))))) True\" apply(simp add: pair_spmf_alt_def R1_14_def R1_14_interm1_def A1'_def Let_def split_def) apply(subst bind_commute_spmf[of "S1_OT12 _ _"]) apply(subst bind_commute_spmf[of "S1_OT12 _ _"]) apply(subst bind_commute_spmf[of "S1_OT12 _ _"]) apply(subst bind_commute_spmf[of "S1_OT12 _ _"]) apply(subst bind_commute_spmf[of "S1_OT12 _ _"]) by auto then show ?thesis by auto qed lemma reduction_step1': shows "\spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\ \ adv_OT12" (is "?lhs \ adv_OT12") proof- have int1: "integrable (measure_spmf (pair_spmf coin_spmf coin_spmf)) (\x. spmf (case x of (m0, m1) \ R1_OT12 (m0, m1) c0 \ (\view. A1 view (m0, m1))) True)" and int2: "integrable (measure_spmf (pair_spmf coin_spmf coin_spmf)) (\x. spmf (case x of (m0, m1) \ S1_OT12 (m0, m1) () \ (\view. A1 view (m0, m1))) True)" by(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1)+ have "?lhs = \LINT x|measure_spmf (pair_spmf coin_spmf coin_spmf). spmf (case x of (m0, m1) \ R1_OT12 (m0, m1) c0 \ (\view. A1 view (m0, m1))) True - spmf (case x of (m0, m1) \ S1_OT12 (m0, m1) () \ (\view. A1 view (m0, m1))) True\" apply(subst (1 2) spmf_bind) using int1 int2 by simp also have "... \ LINT x|measure_spmf (pair_spmf coin_spmf coin_spmf). \spmf (R1_OT12 x c0 \ (\view. A1 view x)) True - spmf (S1_OT12 x () \ (\view. A1 view x)) True\" by(rule integral_abs_bound[THEN order_trans]; simp add: split_beta) ultimately have "?lhs \ LINT x|measure_spmf (pair_spmf coin_spmf coin_spmf). \spmf (R1_OT12 x c0 \ (\view. A1 view x)) True - spmf (S1_OT12 x () \ (\view. A1 view x)) True\" by simp also have "LINT x|measure_spmf (pair_spmf coin_spmf coin_spmf). \spmf (R1_OT12 x c0 \ (\view::'v_OT121. A1 view x)) True - spmf (S1_OT12 x () \ (\view::'v_OT121. A1 view x)) True\ \ adv_OT12" apply(rule integral_mono[THEN order_trans]) apply(rule measure_spmf.integrable_const_bound[where B=2]) apply clarsimp apply(rule abs_triangle_ineq4[THEN order_trans]) subgoal for m0 m1 using pmf_le_1[of "R1_OT12 (m0, m1) c0 \ (\view. A1 view (m0, m1))" "Some True"] pmf_le_1[of "S1_OT12 (m0, m1) () \ (\view. A1 view (m0, m1))" "Some True"] by simp apply simp apply(rule measure_spmf.integrable_const) apply clarify apply(rule OT_12_P1_assms_bound'[rule_format]) by simp ultimately show ?thesis by simp qed lemma reduction_step2: shows "\ A1. \spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\" proof- define A1' where "A1' == \ (view :: 'v_OT121) (m0,m1). do { S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ S1_OT12 (S2,S3) (); c :: 'v_OT121 \ R1_OT12 (S4, S5) c1; let R = (M, (S2,S3, m0, m1, S4, S5), a, view, c); D R}" have "\spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1' view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1' view (m0,m1))))) True\" proof- have "(bind_spmf (R1_14_interm1 M (c0, c1)) D) = (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1' view (m0,m1)))))" unfolding R1_14_interm1_def R1_14_interm2_def A1'_def Let_def split_def apply(simp add: pair_spmf_alt_def) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) including monad_normalisation by(simp) also have "(bind_spmf (R1_14_interm2 M (c0, c1)) D) = (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1' view (m0,m1)))))" unfolding R1_14_interm1_def R1_14_interm2_def A1'_def Let_def split_def apply(simp add: pair_spmf_alt_def) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "_ = \" bind_commute_spmf) by(simp) ultimately show ?thesis by simp qed then show ?thesis by auto qed lemma reduction_step3: shows "\ A1. \spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True - spmf (bind_spmf (S1_14 M out) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\" proof- define A1' where "A1' == \ (view :: 'v_OT121) (m0,m1). do { S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a :: 'v_OT121 \ S1_OT12 (S2,S3) (); b :: 'v_OT121 \ S1_OT12 (S4, S5) (); let R = (M, (S2,S3, S4, S5,m0, m1), a, b, view); D R}" have "\spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True - spmf (bind_spmf (S1_14 M out) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1' view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1' view (m0,m1))))) True\" proof- have "(bind_spmf (R1_14_interm2 M (c0, c1)) D) = (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A1' view (m0,m1)))))" unfolding R1_14_interm2_def A1'_def Let_def split_def apply(simp add: pair_spmf_alt_def) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) including monad_normalisation by(simp) also have "(bind_spmf (S1_14 M out) D) = (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1' view (m0,m1)))))" unfolding S1_14_def Let_def A1'_def split_def apply(simp add: pair_spmf_alt_def) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) including monad_normalisation by(simp) ultimately show ?thesis by simp qed then show ?thesis by auto qed lemma reduction_P1_interm: shows "\spmf (bind_spmf (R1_14 M (c0,c1)) (D)) True - spmf (bind_spmf (S1_14 M out) (D)) True\ \ 3 * adv_OT12" (is "?lhs \ ?rhs") proof- have lhs: "?lhs \ \spmf (bind_spmf (R1_14 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True\ + \spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True\ + \spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True - spmf (bind_spmf (S1_14 M out) D) True\" by simp obtain A1 where A1: "\spmf (bind_spmf (R1_14 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\" using reduction_step1 by blast obtain A2 where A2: "\spmf (bind_spmf (R1_14_interm1 M (c0, c1)) D) True - spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A2 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A2 view (m0,m1))))) True\" using reduction_step2 by blast obtain A3 where A3: "\spmf (bind_spmf (R1_14_interm2 M (c0, c1)) D) True - spmf (bind_spmf (S1_14 M out) D) True\ = \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A3 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A3 view (m0,m1))))) True\" using reduction_step3 by blast have lhs_bound: "?lhs \ \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\ + \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A2 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A2 view (m0,m1))))) True\ + \spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A3 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A3 view (m0,m1))))) True\" using A1 A2 A3 lhs by simp have bound1: "\spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c0) (\ view. (A1 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A1 view (m0,m1))))) True\ \ adv_OT12" and bound2: "\spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A2 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A2 view (m0,m1))))) True\ \ adv_OT12" and bound3: "\spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (R1_OT12 (m0,m1) c1) (\ view. (A3 view (m0,m1))))) True - spmf (bind_spmf (pair_spmf coin_spmf coin_spmf) (\(m0, m1). bind_spmf (S1_OT12 (m0,m1) ()) (\ view. (A3 view (m0,m1))))) True\ \ adv_OT12" using reduction_step1' by auto thus ?thesis using reduction_step1' lhs_bound by argo qed lemma reduction_P1: "\spmf (bind_spmf (R1_14 M (c0,c1)) (D)) True - spmf (funct_OT_14 M (c0,c1) \ (\ (out1,out2). S1_14 M out1 \ (\ view. D view))) True\ \ 3 * adv_OT12" by(simp add: funct_OT_14_def split_def Let_def reduction_P1_interm ) text\Party 2 security.\ lemma coin_coin: "map_spmf (\ S0. S0 \ S3 \ m1) coin_spmf = coin_spmf" (is "?lhs = ?rhs") proof- have lhs: "?lhs = map_spmf (\ S0. S0 \ (S3 \ m1)) coin_spmf" by blast also have op_eq: "... = map_spmf ((\) (S3 \ m1)) coin_spmf" by (metis xor_bool_def) also have "... = ?rhs" using xor_uni_samp by fastforce ultimately show ?thesis using op_eq by auto qed lemma coin_coin': "map_spmf (\ S3. S0 \ S3 \ m1) coin_spmf = coin_spmf" proof- have "map_spmf (\ S3. S0 \ S3 \ m1) coin_spmf = map_spmf (\ S3. S3 \ S0 \ m1) coin_spmf" by (metis xor_left_commute) thus ?thesis using coin_coin by simp qed definition R2_14:: "input1 \ input2 \ 'v_OT122 view2 spmf" where "R2_14 M C = do { let (m0,m1,m2,m3) = M; let (c0,c1) = C; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; let a1 = S0 \ S3 \ m1; let a2 = S1 \ S4 \ m2; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ R2_OT12 (S0,S1) c0; b :: 'v_OT122 \ R2_OT12 (S2,S3) c1; c :: 'v_OT122 \ R2_OT12 (S4,S5) c1; return_spmf (C, (a0,a1,a2,a3), a,b,c)}" lemma lossless_R2_14: "lossless_spmf (R2_14 M C)" by(simp add: R2_14_def split_def lossless_R2_12) definition S2_14 :: "input2 \ bool \ 'v_OT122 view2 spmf" where "S2_14 C out = do { let ((c0::bool),(c1::bool)) = C; S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; a1 :: bool \ coin_spmf; a2 :: bool \ coin_spmf; a3 :: bool \ coin_spmf; let a0' = (if ((\ c0) \ (\ c1)) then (S0 \ S2 \ out) else a0); let a1' = (if ((\ c0) \ c1) then (S0 \ S3 \ out) else a1); let a2' = (if (c0 \ (\ c1)) then (S1 \ S4 \ out) else a2); let a3' = (if (c0 \ c1) then (S1 \ S5 \ out) else a3); a :: 'v_OT122 \ S2_OT12 (c0::bool) (if c0 then S1 else S0); b :: 'v_OT122 \ S2_OT12 (c1::bool) (if c1 then S3 else S2); c :: 'v_OT122 \ S2_OT12 (c1::bool) (if c1 then S5 else S4); return_spmf ((c0,c1), (a0',a1',a2',a3'), a,b,c)}" lemma lossless_S2_14: "lossless_spmf (S2_14 c out)" by(simp add: S2_14_def lossless_S2_12 split_def) lemma P2_OT_14_FT: "R2_14 (m0,m1,m2,m3) (False,True) = funct_OT_14 (m0,m1,m2,m3) (False,True) \ (\ (out1, out2). S2_14 (False,True) out2)" including monad_normalisation proof- have "R2_14 (m0,m1,m2,m3) (False,True) = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ map_spmf (\ S2. S0 \ S2 \ m0) coin_spmf; let a1 = S0 \ S3 \ m1; a2 \ map_spmf (\ S4. S1 \ S4 \ m2) coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((False,True), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def Let_def R2_14_def inf_th_OT12_P2 funct_OT_12_def OT_12_P2_assm) also have "... = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; let a1 = S0 \ S3 \ m1; a2 \ coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((False,True), (a0,a1,a2,a3), a,b,c)}" using coin_coin' by simp also have "... = do { S0 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; let a1 = S0 \ S3 \ m1; a2 :: bool \ coin_spmf; a3 \ map_spmf (\ S1. S1 \ S5 \ m3) coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((False,True), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { S0 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; let a1 = S0 \ S3 \ m1; a2 :: bool \ coin_spmf; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((False,True), (a0,a1,a2,a3), a,b,c)}" using coin_coin by simp ultimately show ?thesis by(simp add: funct_OT_14_def S2_14_def bind_spmf_const) qed lemma P2_OT_14_TT: "R2_14 (m0,m1,m2,m3) (True,True) = funct_OT_14 (m0,m1,m2,m3) (True,True) \ (\ (out1, out2). S2_14 (True,True) out2)" including monad_normalisation proof- have "R2_14 (m0,m1,m2,m3) (True,True) = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ map_spmf (\ S2. S0 \ S2 \ m0) coin_spmf; let a1 = S0 \ S3 \ m1; a2 \ map_spmf (\ S4. S1 \ S4 \ m2) coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((True,True), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def R2_14_def inf_th_OT12_P2 funct_OT_12_def OT_12_P2_assm Let_def) also have "... = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; let a1 = S0 \ S3 \ m1; a2 \ coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((True,True), (a0,a1,a2,a3), a,b,c)}" using coin_coin' by simp also have "... = do { S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; a1 :: bool \ map_spmf (\ S0. S0 \ S3 \ m1) coin_spmf; a2 \ coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((True,True), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { S1 :: bool \ coin_spmf; S3 :: bool \ coin_spmf; S5 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; a1 :: bool \ coin_spmf; a2 \ coin_spmf; let a3 = S1 \ S5 \ m3; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 True S3; c :: 'v_OT122 \ S2_OT12 True S5; return_spmf ((True,True), (a0,a1,a2,a3), a,b,c)}" using coin_coin by simp ultimately show ?thesis by(simp add: funct_OT_14_def S2_14_def bind_spmf_const) qed lemma P2_OT_14_FF: "R2_14 (m0,m1,m2,m3) (False, False) = funct_OT_14 (m0,m1,m2,m3) (False, False) \ (\ (out1, out2). S2_14 (False, False) out2)" including monad_normalisation proof- have "R2_14 (m0,m1,m2,m3) (False,False) = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ map_spmf (\ S3. S0 \ S3 \ m1) coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ map_spmf (\ S5. S1 \ S5 \ m3) coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((False,False), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def R2_14_def inf_th_OT12_P2 funct_OT_12_def OT_12_P2_assm Let_def) also have "... = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((False,False), (a0,a1,a2,a3), a,b,c)}" using coin_coin' by simp also have "... = do { S0 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ coin_spmf; a2 :: bool \ map_spmf (\ S1. S1 \ S4 \ m2) coin_spmf; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((False,False), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { S0 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ coin_spmf; a2 :: bool \ coin_spmf; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 False S0; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((False,False), (a0,a1,a2,a3), a,b,c)}" using coin_coin by simp ultimately show ?thesis by(simp add: funct_OT_14_def S2_14_def bind_spmf_const) qed lemma P2_OT_14_TF: "R2_14 (m0,m1,m2,m3) (True,False) = funct_OT_14 (m0,m1,m2,m3) (True,False) \ (\ (out1, out2). S2_14 (True,False) out2)" including monad_normalisation proof- have "R2_14 (m0,m1,m2,m3) (True,False) = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ map_spmf (\ S3. S0 \ S3 \ m1) coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ map_spmf (\ S5. S1 \ S5 \ m3) coin_spmf; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((True,False), (a0,a1,a2,a3), a,b,c)}" apply(simp add: R2_14_def inf_th_OT12_P2 OT_12_P2_assm funct_OT_12_def Let_def) apply(rewrite in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "\ = _" bind_commute_spmf) by(simp add: bind_map_spmf o_def Let_def) also have "... = do { S0 :: bool \ coin_spmf; S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; let a0 = S0 \ S2 \ m0; a1 :: bool \ coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((True,False), (a0,a1,a2,a3), a,b,c)}" using coin_coin' by simp also have "... = do { S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; a0 :: bool \ map_spmf (\ S0. S0 \ S2 \ m0) coin_spmf; a1 :: bool \ coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((True,False), (a0,a1,a2,a3), a,b,c)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { S1 :: bool \ coin_spmf; S2 :: bool \ coin_spmf; S4 :: bool \ coin_spmf; a0 :: bool \ coin_spmf; a1 :: bool \ coin_spmf; let a2 = S1 \ S4 \ m2; a3 \ coin_spmf; a :: 'v_OT122 \ S2_OT12 True S1; b :: 'v_OT122 \ S2_OT12 False S2; c :: 'v_OT122 \ S2_OT12 False S4; return_spmf ((True,False), (a0,a1,a2,a3), a,b,c)}" using coin_coin by simp ultimately show ?thesis apply(simp add: funct_OT_14_def S2_14_def bind_spmf_const) apply(rewrite in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) apply(rewrite in "bind_spmf _ \" in "bind_spmf _ \" in "bind_spmf _ \" in "_ = \" bind_commute_spmf) by simp qed lemma P2_sec_OT_14_split: "R2_14 (m0,m1,m2,m3) (c0,c1) = funct_OT_14 (m0,m1,m2,m3) (c0,c1) \ (\ (out1, out2). S2_14 (c0,c1) out2)" by(cases c0; cases c1; auto simp add: P2_OT_14_FF P2_OT_14_TF P2_OT_14_FT P2_OT_14_TT) lemma P2_sec_OT_14: "R2_14 M C = funct_OT_14 M C \ (\ (out1, out2). S2_14 C out2)" by(metis P2_sec_OT_14_split surj_pair) sublocale OT_14: sim_det_def R1_14 S1_14 R2_14 S2_14 funct_OT_14 protocol_14_OT unfolding sim_det_def_def by(simp add: lossless_R1_14 lossless_S1_14 lossless_funct_14_OT lossless_R2_14 lossless_S2_14 ) lemma correctness_OT_14: shows "funct_OT_14 M C = protocol_14_OT M C" proof- have "S1 = (S5 = (S1 = (S5 = d))) = d" for S1 S5 d by auto thus ?thesis by(cases "fst C"; cases "snd C"; simp add: funct_OT_14_def protocol_14_OT_def correct funct_OT_12_def lossless_funct_OT_12 bind_spmf_const split_def) qed lemma OT_14_correct: "OT_14.correctness M C" unfolding OT_14.correctness_def using correctness_OT_14 by auto -lemma OT_14_P2_sec: "OT_14.inf_theoretic_P2 m1 m2" - unfolding OT_14.inf_theoretic_P2_def +lemma OT_14_P2_sec: "OT_14.perfect_sec_P2 m1 m2" + unfolding OT_14.perfect_sec_P2_def using P2_sec_OT_14 by blast lemma OT_14_P1_sec: "OT_14.adv_P1 m1 m2 D \ 3 * adv_OT12" unfolding OT_14.adv_P1_def by (metis reduction_P1 surj_pair) end locale OT_14_asymp = sim_det_def + fixes S1_OT12 :: "nat \ (bool \ bool) \ unit \ 'v_OT121 spmf" and R1_OT12 :: "nat \ (bool \ bool) \ bool \ 'v_OT121 spmf" and adv_OT12 :: "nat \ real" and S2_OT12 :: "nat \ bool \ bool \ 'v_OT122 spmf" and R2_OT12 :: "nat \ (bool \ bool) \ bool \ 'v_OT122 spmf" and protocol_OT12 :: "(bool \ bool) \ bool \ (unit \ bool) spmf" assumes ot14_base: "\ (n::nat). ot14_base (S1_OT12 n) (R1_12_0T n) (adv_OT12 n) (S2_OT12 n) (R2_12OT n) (protocol_OT12)" begin sublocale ot14_base "(S1_OT12 n)" "(R1_12_0T n)" "(adv_OT12 n)" "(S2_OT12 n)" "(R2_12OT n)" using local.ot14_base by simp lemma OT_14_P1_sec: "OT_14.adv_P1 (R1_12_0T n) n m1 m2 D \ 3 * (adv_OT12 n)" unfolding OT_14.adv_P1_def using reduction_P1 surj_pair by metis theorem OT_14_P1_asym_sec: "negligible (\ n. OT_14.adv_P1 (R1_12_0T n) n m1 m2 D)" if "negligible (\ n. adv_OT12 n)" proof- have adv_neg: "negligible (\n. 3 * adv_OT12 n)" using that negligible_cmultI by simp have "\OT_14.adv_P1 (R1_12_0T n) n m1 m2 D\ \ \3 * (adv_OT12 n)\" for n proof - have "\OT_14.adv_P1 (R1_12_0T n) n m1 m2 D\ \ 3 * adv_OT12 n" using OT_14.adv_P1_def OT_14_P1_sec by auto then show ?thesis by (meson abs_ge_self order_trans) qed thus ?thesis using OT_14_P1_sec negligible_le adv_neg by (metis (no_types, lifting) negligible_absI) qed -theorem OT_14_P2_asym_sec: "OT_14.inf_theoretic_P2 R2_OT12 n m1 m2" +theorem OT_14_P2_asym_sec: "OT_14.perfect_sec_P2 R2_OT12 n m1 m2" using OT_14_P2_sec by simp end end diff --git a/thys/Multi_Party_Computation/Secure_Multiplication.thy b/thys/Multi_Party_Computation/Secure_Multiplication.thy --- a/thys/Multi_Party_Computation/Secure_Multiplication.thy +++ b/thys/Multi_Party_Computation/Secure_Multiplication.thy @@ -1,800 +1,800 @@ subsection \Secure multiplication protocol\ theory Secure_Multiplication imports CryptHOL.Cyclic_Group_SPMF Uniform_Sampling Semi_Honest_Def begin locale secure_mult = fixes q :: nat assumes q_gt_0: "q > 0" and "prime q" begin type_synonym real_view = "nat \ nat \ ((nat \ nat \ nat \ nat) \ nat \ nat) spmf" type_synonym sim = "nat \ nat \ ((nat \ nat \ nat \ nat) \ nat \ nat) spmf" lemma samp_uni_set_spmf [simp]: "set_spmf (sample_uniform q) = {..< q}" by(simp add: sample_uniform_def) definition funct :: "nat \ nat \ (nat \ nat) spmf" where "funct x y = do { s \ sample_uniform q; return_spmf (s, (x*y + (q - s)) mod q)}" definition TI :: "((nat \ nat) \ (nat \ nat)) spmf" where "TI = do { a \ sample_uniform q; b \ sample_uniform q; r \ sample_uniform q; return_spmf ((a, r), (b, ((a*b + (q - r)) mod q)))}" definition out :: "nat \ nat \ (nat \ nat) spmf" where "out x y = do { ((c1,d1),(c2,d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; return_spmf (((x*e1 + (q - d1)) mod q), ((e2 * c2 + (q - d2)) mod q))}" definition R1 :: "real_view" where "R1 x y = do { ((c1, d1), (c2, d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((x, c1, d1, e1), s1, s2)}" definition S1 :: "nat \ nat \ (nat \ nat \ nat \ nat) spmf" where "S1 x s1 = do { a :: nat \ sample_uniform q; e1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf (x, a, d1, e1)}" definition Out1 :: "nat \ nat \ nat \ (nat \ nat) spmf" where "Out1 x y s1 = do { let s2 = (x*y + (q - s1)) mod q; return_spmf (s1,s2)}" definition R2 :: "real_view" where "R2 x y = do { ((c1, d1), (c2, d2)) \ TI; let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((y, c2, d2, e2), s1, s2)}" definition S2 :: "nat \ nat \ (nat \ nat \ nat \ nat) spmf" where "S2 y s2 = do { b \ sample_uniform q; e2 \ sample_uniform q; let d2 = (e2*b + (q - s2)) mod q; return_spmf (y, b, d2, e2)}" definition Out2 :: "nat \ nat \ nat \ (nat \ nat) spmf" where "Out2 y x s2 = do { let s1 = (x*y + (q - s2)) mod q; return_spmf (s1,s2)}" definition Ideal2 :: "nat \ nat \ nat \ ((nat \ nat \ nat \ nat) \ (nat \ nat)) spmf" where "Ideal2 y x out2 = do { view2 :: (nat \ nat \ nat \ nat) \ S2 y out2; out2 \ Out2 y x out2; return_spmf (view2, out2)}" sublocale sim_non_det_def: sim_non_det_def R1 S1 Out1 R2 S2 Out2 funct . lemma minus_mod: assumes "a > b" shows "[a - b mod q = a - b] (mod q)" by(metis cong_diff_nat cong_def le_trans less_or_eq_imp_le assms mod_less_eq_dividend mod_mod_trivial) lemma q_cong:"[a = q + a] (mod q)" by (simp add: cong_def) lemma q_cong_reverse: "[q + a = a] (mod q)" by (simp add: cong_def) lemma qq_cong: "[a = q*q + a] (mod q)" by (simp add: cong_def) lemma minus_q_mult_cancel: assumes "[a = e + b - q * c - d] (mod q)" and "e + b - d > 0" and "e + b - q * c - d > 0" shows "[a = e + b - d] (mod q)" proof- have "a mod q = (e + b - q * c - d) mod q" using assms(1) cong_def by blast then have "a mod q = (e + b - d) mod q" by (metis (no_types) add_cancel_left_left assms(3) diff_commute diff_is_0_eq' linordered_semidom_class.add_diff_inverse mod_add_left_eq mod_mult_self1_is_0 nat_less_le) then show ?thesis using cong_def by blast qed lemma s1_s2: assumes "x < q" "a < q" "b < q" and r:"r < q" "y < q" shows "((x + a) mod q * b + q - (a * b + q - r) mod q) mod q = (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q" proof- have s: "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q = ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" proof- have lhs: "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q = (x*b + r) mod q" proof- let ?h = "(x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q" have "[?h = x * y + q - (x * ((y + (q - b)) mod q) + (q - r)) mod q] (mod q)" by(simp add: assms(1) cong_def q_gt_0) then have "[?h = x * y + q - (x * (y + (q - b)) + (q - r)) mod q] (mod q)" by (metis mod_add_left_eq mod_mult_right_eq) then have no_qq: "[?h = x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by(metis distrib_left) then have "[?h = q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" proof- have "[x * y + q - (x * y + x * (q - b) + (q - r)) mod q = q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by (smt qq_cong add.assoc cong_diff_nat cong_def le_add2 le_trans mod_le_divisor q_gt_0) then show ?thesis using cong_trans no_qq by blast qed then have mod: "[?h = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def add.assoc add.commute le_add2 le_trans mod_add_self2 mod_le_divisor q_gt_0) then have "[?h = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r))] (mod q)" proof- have 1: "q \ q - b" using assms by simp then have "q*q \ x*(q-b)" "q \ q - r" using 1 assms apply (auto simp add: mult_strict_mono) by (simp add: mult_le_mono) then have "q + q*q + x * y + q > x * y + x * (q - b) + (q - r)" using assms(5) by linarith then have "[q + q*q + x * y + q - (x * y + x * (q - b) + (q - r)) mod q = q + q*q + x * y + q - (x * y + x * (q - b) + (q - r))] (mod q)" using minus_mod by blast then show ?thesis using mod using cong_trans by blast qed then have "[?h = q + q*q + x * y + q - (x * y + (x * q - x*b) + (q - r))] (mod q)" by (simp add: right_diff_distrib') then have "[?h = q + q*q + x * y + q - x * y - (x * q - x*b) - (q - r)] (mod q)" by simp then have mod': "[?h = q + q*q + q - (x * q - x*b) - (q - r)] (mod q)" by(simp add: add.commute) then have neg: "[?h = q + q*q + q - x * q + x*b - (q - r)] (mod q)" proof- have "[q + q*q + q - (x * q - x*b) - (q - r) = q + q*q + q - x * q + x*b - (q - r)] (mod q)" proof(cases "x = 0") case True then show ?thesis by simp next case False have "x * q - x*b > 0" using False assms by simp also have "q + q*q + q - x * q > 0" by (metis assms(1) add.commute diff_mult_distrib2 less_Suc_eq mult.commute mult_Suc_right nat_0_less_mult_iff q_gt_0 zero_less_diff) ultimately show ?thesis by simp qed then show ?thesis using mod' cong_trans by blast qed then have "[?h = q + q*q + q + x*b - (q - r)] (mod q)" proof- have "[q + q*q + q - x * q + x*b - (q - r) = q + q*q + q + x*b - (q - r)] (mod q)" proof(cases "x = 0") case True then show ?thesis by simp next case False have "q*q > x*q" using False assms by (simp add: mult_strict_mono) then have 1: "q + q*q + q - x * q + x*b - (q - r) > 0" by linarith then have 2: "q + q*q + q + x*b - (q - r) > 0" by simp then show ?thesis by (smt 1 2 Nat.add_diff_assoc2 \x * q < q * q\ add_cancel_left_left add_diff_inverse_nat le_add1 le_add2 le_trans less_imp_add_positive less_numeral_extra(3) minus_mod minus_q_mult_cancel mod_if mult.commute q_gt_0) qed then show ?thesis using cong_trans neg by blast qed then have "[?h = q + q*q + q + x*b - q + r] (mod q)" by (metis r(1) Nat.add_diff_assoc2 Nat.diff_diff_right le_add2 less_imp_le_nat semiring_normalization_rules(23)) then have "[?h = q + q*q + q + x*b + r] (mod q)" apply(simp add: cong_def) by (metis (no_types, lifting) add.assoc add.commute add_diff_cancel_right' diff_is_0_eq' mod_if mod_le_divisor q_gt_0) then have "[?h = x*b + r] (mod q)" apply(simp add: cong_def) by (metis mod_add_cong mod_add_self1 mod_mult_self1) then show ?thesis by (simp add: cong_def assms) qed also have rhs: "((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q = (x*b + r) mod q" proof- let ?h = "((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" have "[?h = (x + a) mod q * b + q - (a * b + (q - r)) mod q] (mod q)" by (simp add: q_gt_0 assms(1) cong_def) then have "[?h = (x + a) * b + q - (a * b + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def mod_add_cong mod_le_divisor mod_mult_left_eq q_gt_0 assms) then have "[?h = x*b + a*b + q - (a * b + (q - r)) mod q] (mod q)" by(metis distrib_right) then have mod: "[?h = q + x*b + a*b + q - (a * b + (q - r)) mod q] (mod q)" by (smt Nat.add_diff_assoc cong_def add.assoc add.commute le_add2 le_trans mod_add_self2 mod_le_divisor q_gt_0) then have "[?h = q + x*b + a*b + q - (a * b + (q - r))] (mod q)" using q_cong assms(1) proof- have ge: "q + x*b + a*b + q > (a * b + (q - r))" using assms by simp then have "[ q + x*b + a*b + q - (a * b + (q - r)) mod q = q + x*b + a*b + q - (a * b + (q - r))] (mod q)" using Divides.mod_less_eq_dividend cong_diff_nat cong_def le_trans less_not_refl2 less_or_eq_imp_le q_gt_0 minus_mod by presburger then show ?thesis using mod cong_trans by blast qed then have "[?h = q + x*b + q - (q - r)] (mod q)" by (simp add: add.commute) then have "[?h = q + x*b + q - q + r] (mod q)" by (metis Nat.add_diff_assoc2 Nat.diff_diff_right r(1) le_add2 less_imp_le_nat) then have "[?h = q + x*b + r] (mod q)" by simp then have "[?h = q + (x*b + r)] (mod q)" using add.assoc by metis then have "[?h = x*b + r] (mod q)" using cong_def q_cong_reverse by metis then show ?thesis by (simp add: cong_def assms(1)) qed ultimately show ?thesis by simp qed have lhs:"((x + a) mod q * b + q - (a * b + q - r) mod q) mod q = ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q" using assms by simp have rhs: "(x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q = (x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q " using assms by simp have " ((x + a) mod q * b + (q - (a * b + (q - r)) mod q)) mod q = (x * y + (q - (x * ((y + (q - b)) mod q) + (q - r)) mod q)) mod q" using assms s[symmetric] by blast then show ?thesis using lhs rhs by metis qed lemma s1_s2_P2: assumes "x < q" "xa < q" "xb < q" "xc < q" "y < q" shows "((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, ((x + xb) mod q * xa + q - (xb * xa + q - xc) mod q) mod q) = ((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q)" using assms s1_s2 by metis lemma c1: assumes "e2 = (x + c1) mod q" and "x < q" "c1 < q" shows "c1 = (e2 + q - x) mod q" proof- have "[e2 + q = x + c1] (mod q)" by(simp add: assms cong_def) then have "[e2 + q - x = c1] (mod q)" proof- have "e2 + q \ x" using assms by simp then show ?thesis by (metis \[e2 + q = x + c1] (mod q)\ cong_add_lcancel_nat le_add_diff_inverse) qed then show ?thesis by(simp add: cong_def assms) qed lemma c1_P2: assumes "xb < q" "xa < q" "xc < q" "x < q" shows "((y, xa, (xb * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q) = ((y, xa, (((x + xb) mod q + q - x) mod q * xa + q - xc) mod q, (x + xb) mod q), (x * ((y + q - xa) mod q) + q - xc) mod q, (x * y + q - (x * ((y + q - xa) mod q) + q - xc) mod q) mod q)" proof- have "(xb * xa + q - xc) mod q = (((x + xb) mod q + q - x) mod q * xa + q - xc) mod q" using assms c1 by simp then show ?thesis using assms by metis qed lemma minus_mod_cancel: assumes "a - b > 0" "a - b mod q > 0" shows "[a - b + c = a - b mod q + c] (mod q)" proof- have "(b - b mod q + (a - b)) mod q = (0 + (a - b)) mod q" using cong_def mod_add_cong neq0_conv q_gt_0 by (simp add: minus_mod_eq_mult_div) then show ?thesis by (metis (no_types) Divides.mod_less_eq_dividend Nat.add_diff_assoc2 add_diff_inverse_nat assms(1) cong_def diff_is_0_eq' less_or_eq_imp_le mod_add_cong neq0_conv) qed lemma d2: assumes d2: "d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q" and s1: "s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q" and s2: "s2 = (x*y + (q - s1)) mod q" and x: "x < q" and y: "y < q" and r: "r < q" and b: "b < q" and e2: "e2 < q" shows "d2 = (e2*b + (q - s2)) mod q" proof- have s1_le_q: "s1 < q" using s1 q_gt_0 by simp have s2_le_q: "s2 < q" using s2 q_gt_0 by simp have xb: "(x*b) mod q = (s2 + (q - r)) mod q" proof- have "s1 = (x*(y + (q - b)) + (q - r)) mod q" using s1 b by (metis mod_add_left_eq mod_mult_right_eq) then have s1_dist: "s1 = (x*y + x*(q - b) + (q - r)) mod q" by(metis distrib_left) then have "s1 = (x*y + x*q - x*b + (q - r)) mod q" using distrib_left b diff_mult_distrib2 by auto then have "[s1 = x*y + x*q - x*b + (q - r)] (mod q)" by(simp add: cong_def) then have "[s1 + x * b = x*y + x*q - x*b + x*b + (q - r)] (mod q)" by (metis add.commute add.left_commute cong_add_lcancel_nat) then have "[s1 + x*b = x*y + x*q + (q - r)] (mod q)" using b by (simp add: algebra_simps) (metis add_diff_inverse_nat diff_diff_left diff_mult_distrib2 less_imp_add_positive mult.commute not_add_less1 zero_less_diff) then have s1_xb: "[s1 + x*b = q + x*y + x*q + (q - r)] (mod q)" by (smt mod_add_cong mod_add_self1 cong_def) then have "[x*b = q + x*y + x*q + (q - r) - s1] (mod q)" proof- have "q + x*y + x*q + (q - r) - s1 > 0" using s1_le_q by simp then show ?thesis by (metis add_diff_inverse_nat less_numeral_extra(3) s1_xb cong_add_lcancel_nat nat_diff_split) qed then have "[x*b = x*y + x*q + (q - r) + q - s1] (mod q)" by (metis add.assoc add.commute) then have "[x*b = x*y + (q - r) + q - s1] (mod q)" by (smt Nat.add_diff_assoc cong_def less_imp_le_nat mod_mult_self1 s1_le_q semiring_normalization_rules(23)) then have "(x*b) mod q = (x*y + (q - r) + q - s1) mod q" by(simp add: cong_def) then have "(x*b) mod q = (x*y + (q - r) + (q - s1)) mod q" using add.assoc s1_le_q by auto then have "(x*b) mod q = (x*y + (q - s1) + (q - r)) mod q" using add.commute by presburger then show ?thesis using s2 by presburger qed have "d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q" using d2 by simp then have "d2 = (((e2 + q - x))*b + (q - r)) mod q" using mod_add_cong mod_mult_left_eq by blast then have "d2 = (e2*b + q*b - x*b + (q - r)) mod q" by (simp add: distrib_right diff_mult_distrib) then have a: "[d2 = e2*b + q*b - x*b + (q - r)] (mod q)" by(simp add: cong_def) then have b:"[d2 = q + q + e2*b + q*b - x*b + (q - r)] (mod q)" proof- have "[e2*b + q*b - x*b + (q - r) = q + q + e2*b + q*b - x*b + (q - r)] (mod q)" by (smt assms Nat.add_diff_assoc add.commute cong_def less_imp_le_nat mod_add_self2 mult.commute nat_mult_le_cancel_disj semiring_normalization_rules(23)) then show ?thesis using cong_trans a by blast qed then have "[d2 = q + q + e2*b + q*b - (x*b) mod q + (q - r)] (mod q)" proof- have "[q + q + e2*b + q*b - (x*b) + (q - r) = q + q + e2*b + q*b - (x*b) mod q + (q - r)] (mod q)" proof(cases "b = 0") case True then show ?thesis by simp next case False have "q*b - (x*b) > 0" using False x by simp then have 1: "q + q + e2*b + q*b - (x*b) > 0" by linarith then have 2:"q + q + e2*b + q*b - (x*b) mod q > 0" by (simp add: q_gt_0 trans_less_add1) then show ?thesis using 1 2 minus_mod_cancel by simp qed then show ?thesis using cong_trans b by blast qed then have c: "[d2 = q + q + e2*b + q*b - (s2 + (q - r)) mod q + (q - r)] (mod q)" using xb by simp then have "[d2 = q + q + e2*b + q*b - (s2 + (q - r)) + (q - r)] (mod q)" proof- have "[q + q + e2*b + q*b - (s2 + (q - r)) mod q + (q - r) = q + q + e2*b + q*b - (s2 + (q - r)) + (q - r)] (mod q)" proof- have "q + q + e2*b + q*b - (s2 + (q - r)) mod q > 0" by (metis diff_is_0_eq gr0I le_less_trans mod_less_divisor not_add_less1 q_gt_0 semiring_normalization_rules(23) trans_less_add2) moreover have"q + q + e2*b + q*b - (s2 + (q - r)) > 0" using s2_le_q by simp ultimately show ?thesis using minus_mod_cancel cong_sym by blast qed then show ?thesis using cong_trans c by blast qed then have d: "[d2 = q + q + e2*b + q*b - s2 - (q - r) + (q - r)] (mod q)" by simp then have "[d2 = q + q + e2*b + q*b - s2 ] (mod q)" proof- have "q + q + e2*b + q*b - s2 - (q - r) > 0" using s2_le_q by simp then show ?thesis using d cong_trans by simp qed then have "[d2 = q + q + e2*b - s2] (mod q)" by (smt Nat.add_diff_assoc2 cong_def less_imp_le_nat mod_mult_self1 mult.commute s2_le_q semiring_normalization_rules(23) trans_less_add2) then have "[d2 = q + e2*b + q - s2] (mod q)" by(simp add: add.commute add.assoc) then have "[d2 = e2*b + q - s2] (mod q)" by (smt Nat.add_diff_assoc2 add.commute cong_def less_imp_le_nat mod_add_self2 s2_le_q trans_less_add2) then have "[d2 = e2*b + (q - s2)] (mod q)" by (simp add: less_imp_le_nat s2_le_q) then show ?thesis by(simp add: cong_def d2) qed lemma d2_P2: assumes x: "x < q" and y: "y < q" and r: "b < q" and b: "e2 < q" and e2: "r < q" shows "((y, b, ((e2 + q - x) mod q * b + q - r) mod q, e2), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) = ((y, b, (e2 * b + q - (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) mod q, e2), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q)" proof- have "((e2 + q - x) mod q * b + q - r) mod q = (e2 * b + q - (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) mod q" (is "?lhs = ?rhs") proof- have d2: "(((e2 + q - x) mod q)*b + (q - r)) mod q = (e2*b + (q - ((x*y + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q))) mod q" using assms d2 by blast have "?lhs = (((e2 + q - x) mod q)*b + (q - r)) mod q" using assms by simp also have "?rhs = (e2*b + (q - ((x*y + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q))) mod q" using assms by simp ultimately show ?thesis using assms d2 by metis qed then show ?thesis using assms by metis qed lemma s1: assumes s2: "s2 = (x*y + q - s1) mod q" and x: "x < q" and y: "y < q" and s1: "s1 < q" shows "s1 = (x*y + q - s2) mod q" proof- have s2_le_q:"s2 < q" using s2 q_gt_0 by simp have "[s2 = x*y + q - s1] (mod q)" by(simp add: cong_def s2) then have "[s2 = x*y + q - s1] (mod q)" using add.assoc by (simp add: less_imp_le_nat s1) then have s1_s2: "[s2 + s1 = x*y + q] (mod q)" by (metis (mono_tags, lifting) cong_def le_add2 le_add_diff_inverse2 le_trans mod_add_left_eq order.strict_implies_order s1) then have "[s1 = x*y + q - s2] (mod q)" proof- have "x*y + q - s2 > 0" using s2_le_q by simp then show ?thesis by (metis s1_s2 add_diff_cancel_left' cong_diff_nat cong_def le_add1 less_imp_le_nat zero_less_diff) qed then show ?thesis by(simp add: cong_def s1) qed lemma s1_P2: assumes x: "x < q" and y: "y < q" and "b < q" and "e2 < q" and "r < q" and "s1 < q" shows "((y, b, (e2 * b + q - (x * y + q - r) mod q) mod q, e2), r, (x * y + q - r) mod q) = ((y, b, (e2 * b + q - (x * y + q - r) mod q) mod q, e2), (x * y + q - (x * y + q - r) mod q) mod q, (x * y + q - r) mod q)" proof- have "s1 = (x*y + q - ((x*y + q - s1) mod q)) mod q" using assms secure_mult.s1 secure_mult_axioms by blast then show ?thesis using assms s1 by blast qed theorem P2_security: assumes "x < q" "y < q" - shows "sim_non_det_def.inf_theoretic_P2 x y" + shows "sim_non_det_def.perfect_sec_P2 x y" including monad_normalisation proof- have "((funct x y) \ (\ (s1',s2'). (sim_non_det_def.Ideal2 y x s2'))) = R2 x y" proof- have "R2 x y = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let d2 = ((a*b + (q - r)) mod q); let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - r)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((y, c2, d2, e2), s1, s2)}" by(simp add: R2_def TI_def Let_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let e2 = (x + c1) mod q; let d2 = ((((e2 + q - x) mod q)*b + (q - r)) mod q); let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), (s1, (x*y + (q - s1)) mod q))}" by(simp add: Let_def s1_s2_P2 assms c1_P2 cong: bind_spmf_cong_simp) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let d1 = r; let c2 = b; e2 \ map_spmf (\ c1. (x + c1) mod q) (sample_uniform q); let d2 = ((((e2 + q - x) mod q)*b + (q - r)) mod q); let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let d1 = r; let c2 = b; e2 \ sample_uniform q; let d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q; let s1 = (x*((y + (q - c2)) mod q) + (q - r)) mod q; return_spmf ((y, c2, d2, e2), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_plus_one_time_pad) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; e2 \ sample_uniform q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let s2 = (x*y + (q - s1)) mod q; let d2 = (((e2 + q - x) mod q)*b + (q - r)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp) also have "... = do { b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; e2 \ sample_uniform q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: d2_P2 assms Let_def cong: bind_spmf_cong_simp) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ map_spmf (\ r. (x*((y + (q - b)) mod q) + (q - r)) mod q) (sample_uniform q); let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: bind_map_spmf o_def Let_def) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ sample_uniform q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), s1, s2)}" by(simp add: samp_uni_minus_one_time_pad) also have "... = do { b :: nat \ sample_uniform q; e2 \ sample_uniform q; s1 \ sample_uniform q; let s2 = (x*y + (q - s1)) mod q; let d2 = (e2*b + (q - s2)) mod q; return_spmf ((y, b, d2, e2), (x*y + (q - s2)) mod q, s2)}" by(simp add: s1_P2 assms Let_def cong: bind_spmf_cong_simp) ultimately show ?thesis by(simp add: funct_def Let_def sim_non_det_def.Ideal2_def Out2_def S2_def R2_def) qed - then show ?thesis by(simp add: sim_non_det_def.inf_theoretic_P2_def) + then show ?thesis by(simp add: sim_non_det_def.perfect_sec_P2_def) qed lemma s1_s2_P1: assumes "x < q" "xa < q" "xb < q" "xc < q" "y < q" shows "((x, xa, xb, (y + q - xc) mod q), (x * ((y + q - xc) mod q) + q - xb) mod q, ((x + xa) mod q * xc + q - (xa * xc + q - xb) mod q) mod q) = ((x, xa, xb, (y + q - xc) mod q), (x * ((y + q - xc) mod q) + q - xb) mod q, (x * y + q - (x * ((y + q - xc) mod q) + q - xb) mod q) mod q)" using assms s1_s2 by metis lemma mod_minus: assumes "a - b > 0" and "c - d > 0" shows "(a - b + (c - d mod q)) mod q = (a - b + (c - d)) mod q" using assms by (metis cong_def minus_mod mod_add_right_eq zero_less_diff) lemma r: assumes e1: "e1 = (y + (q - b)) mod q" and s1: "s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q" and b: "b < q" and x: "x < q" and y: "y < q" and r: "r < q" shows "r = (x*e1 + (q - s1)) mod q" (is "?lhs = ?rhs") proof- have "s1 = (x*((y + (q - b))) + (q - r)) mod q" using s1 b by (metis mod_add_left_eq mod_mult_right_eq) then have s1_dist: "s1 = (x*y + x*(q - b) + (q - r)) mod q" by(metis distrib_left) then have "?rhs = (x*((y + (q - b)) mod q) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" using e1 by simp then have "?rhs = (x*((y + (q - b))) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" by (metis mod_add_left_eq mod_mult_right_eq) then have "?rhs = (x*y + x*(q - b) + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" by(metis distrib_left) then have a: "?rhs = (x*y + x*q - x*b + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" using distrib_left b diff_mult_distrib2 by auto then have b: "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" proof - have "(x*y + x*q - x*b + (q - (x*y + x*(q - b) + (q - r)) mod q)) mod q = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q" proof - have f1: "\n na nb nc nd. (n::nat) mod na \ nb mod na \ nc mod na \ nd mod na \ (n + nc) mod na = (nb + nd) mod na" by (meson mod_add_cong) then have "(q - (x * y + x*(q - b) + (q - r)) mod q) mod q = (q * q + q * q + q - (x * y + x*(q - b) + (q - r)) mod q) mod q" by (metis Nat.diff_add_assoc mod_le_divisor q_gt_0 mod_mult_self4) then show ?thesis using f1 by blast qed then show ?thesis using a by simp qed then have "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)))) mod q" proof- have "(x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b) + (q - r)) mod q)) mod q = (x*y + x*q - x*b + (q*q + q*q + q - (x*y + x*(q - b)+ (q - r)))) mod q" proof(cases "x = 0") case True then show ?thesis by (metis (no_types, lifting) assms(2) assms(4) True Nat.add_diff_assoc add.left_neutral cong_def diff_le_self minus_mod mult_is_0 not_gr_zero zero_eq_add_iff_both_eq_0 zero_less_diff) next case False have qb: "q - b \ q" using b by simp then have qb':"x*(q - b) < q*q" using x by (metis mult_less_le_imp_less nat_0_less_mult_iff nat_less_le neq0_conv) have a: "x*y + x*(q - b) > 0" using False assms by simp have 1: "q*q > x*y" using False by (simp add: mult_strict_mono q_gt_0 x y) have 2: "q*q > x*q" using False by (simp add: mult_strict_mono q_gt_0 x y) have b: "(q*q + q*q + q - (x*y + x*(q - b) + (q - r))) > 0" using 1 qb' by simp then show ?thesis using a b mod_minus[of "x*y + x*q" "x*b" "q*q + q*q + q" "x*y + x*(q - b) + (q - r)"] by (smt add.left_neutral cong_def gr0I minus_mod zero_less_diff) qed then show ?thesis using b by simp qed then have d: "?rhs = (x*y + x*q - x*b + (q*q + q*q + q - x*y - x*(q - b) - (q - r))) mod q" by simp then have e: "?rhs = (x*y + x*q - x*b + q*q + q*q + q - x*y - x*(q - b) - (q - r)) mod q" proof(cases "x = 0") case True then show ?thesis using True d by simp next case False have qb: "q - b \ q" using b by simp then have qb':"x*(q - b) < q*q" using x by (metis mult_less_le_imp_less nat_0_less_mult_iff nat_less_le neq0_conv) have a: "x*y + x*(q - b) > 0" using False assms by simp have 1: "q*q > x*y" using False by (simp add: mult_strict_mono q_gt_0 x y) have 2: "q*q > x*q" using False by (simp add: mult_strict_mono q_gt_0 x y) have b: "q*q + q*q + q - x*y - x*(q - b) - (q - r) > 0" using 1 qb' by simp then show ?thesis using b d by (smt Nat.add_diff_assoc add.assoc diff_diff_left less_imp_le_nat zero_less_diff) qed then have "?rhs = (x*q - x*b + q*q + q*q + q - x*(q - b) - (q - r)) mod q" proof- have "(x*y + x*q - x*b + q*q + q*q + q - x*y - x*(q - b) - (q - r)) mod q = (x*q - x*b + q*q + q*q + q - x*(q - b) - (q - r)) mod q" proof - have 1: "q + n - b = q - b + n" for n by (simp add: assms(3) less_imp_le) have 2: "(c::nat) * b + (c * a + n) = c * (b + a) + n" for n a b c by (simp add: distrib_left) have "(c::nat) + (b + a) - (n + a) = c + b - n" for n a b c by auto then have "(q + (q * q + (q * q + x * (q + y - b))) - (q - r + x * (y + (q - b)))) mod q = (q + (q * q + (q * q + x * (q - b))) - (q - r + x * (q - b))) mod q" by (metis (no_types) add.commute 1 2) then show ?thesis by (simp add: add.commute diff_mult_distrib2 distrib_left) qed then show ?thesis using e by simp qed then have "?rhs = (x*(q - b) + q*q + q*q + q - x*(q - b) - (q - r)) mod q" by(metis diff_mult_distrib2) then have "?rhs = (q*q + q*q + q - (q - r)) mod q" using assms(6) by simp then have "?rhs = (q*q + q*q + q - q + r) mod q" using assms(6) by(simp add: Nat.diff_add_assoc2 less_imp_le) then have "?rhs = (q*q + q*q + r) mod q" by simp then have "?rhs = r mod q" by (metis add.commute distrib_right mod_mult_self1) then show ?thesis using assms(6) by simp qed lemma r_P2: assumes b: "b < q" and x: "x < q" and y: "y < q" and r: "r < q" shows "((x, a, r, (y + q - b) mod q), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q) = ((x, a, (x * ((y + q - b) mod q) + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q, (y + q - b) mod q), (x * ((y + q - b) mod q) + q - r) mod q, (x * y + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q)" proof- have "(x * ((y + q - b) mod q) + q - (x * ((y + q - b) mod q) + q - r) mod q) mod q = r" (is "?lhs = ?rhs") proof- have 1:"r = (x*((y + (q - b)) mod q) + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q" using assms secure_mult.r secure_mult_axioms by blast also have "?rhs = (x*((y + (q - b)) mod q) + (q - ((x*((y + (q - b)) mod q) + (q - r)) mod q))) mod q" using assms 1 by blast ultimately show ?thesis using assms 1 by simp qed then show ?thesis using assms by simp qed theorem P1_security: assumes "x < q" "y < q" - shows "sim_non_det_def.inf_theoretic_P1 x y" + shows "sim_non_det_def.perfect_sec_P1 x y" including monad_normalisation proof- have "(funct x y) \ (\ (s1',s2'). (sim_non_det_def.Ideal1 x y s1')) = R1 x y" proof- have "R1 x y = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let d1 = r; let c2 = b; let d2 = ((a*b + (q - r)) mod q); let e2 = (x + c1) mod q; let e1 = (y + (q - c2)) mod q; let s1 = (x*e1 + (q - d1)) mod q; let s2 = (e2 * c2 + (q - d2)) mod q; return_spmf ((x, c1, d1, e1), s1, s2)}" by(simp add: R1_def TI_def Let_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; r :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; let s1 = (x*((y + (q - b)) mod q) + (q - r)) mod q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: Let_def assms s1_s2_P1 r_P2 cong: bind_spmf_cong_simp) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; s1 \ map_spmf (\ r. (x*((y + (q - b)) mod q) + (q - r)) mod q) (sample_uniform q); let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf Let_def o_def) also have "... = do { a :: nat \ sample_uniform q; b :: nat \ sample_uniform q; let c1 = a; let c2 = b; let e1 = (y + (q - b)) mod q; s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_minus_one_time_pad) also have "... = do { a :: nat \ sample_uniform q; let c1 = a; e1 \ map_spmf (\b. (y + (q - b)) mod q) (sample_uniform q); s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: bind_map_spmf Let_def o_def) also have "... = do { a :: nat \ sample_uniform q; let c1 = a; e1 \ sample_uniform q; s1 \ sample_uniform q; let d1 = (x*e1 + (q - s1)) mod q; return_spmf ((x, c1, d1, e1), s1, (x*y + (q - s1)) mod q)}" by(simp add: samp_uni_minus_one_time_pad) ultimately show ?thesis by(simp add: funct_def sim_non_det_def.Ideal1_def Let_def R1_def TI_def Out1_def S1_def) qed - thus ?thesis by(simp add: sim_non_det_def.inf_theoretic_P1_def) + thus ?thesis by(simp add: sim_non_det_def.perfect_sec_P1_def) qed end locale secure_mult_asymp = fixes q :: "nat \ nat" assumes "\ n. secure_mult (q n)" begin sublocale secure_mult "q n" for n using secure_mult_asymp_axioms secure_mult_asymp_def by blast theorem P1_secure: assumes "x < q n" "y < q n" - shows "sim_non_det_def.inf_theoretic_P1 n x y" + shows "sim_non_det_def.perfect_sec_P1 n x y" by (metis P1_security assms) theorem P2_secure: assumes "x < q n" "y < q n" - shows "sim_non_det_def.inf_theoretic_P2 n x y" + shows "sim_non_det_def.perfect_sec_P2 n x y" by (metis P2_security assms) end end \ No newline at end of file diff --git a/thys/Multi_Party_Computation/Semi_Honest_Def.thy b/thys/Multi_Party_Computation/Semi_Honest_Def.thy --- a/thys/Multi_Party_Computation/Semi_Honest_Def.thy +++ b/thys/Multi_Party_Computation/Semi_Honest_Def.thy @@ -1,186 +1,186 @@ section \Semi-Honest Security\ text \We follow the security definitions for the semi honest setting as described in \cite{DBLP:books/sp/17/Lindell17}. In the semi honest model the parties are assumed not to deviate from the protocol transcript. Semi honest security guarantees that no information is leaked during the running of the protocol.\ subsection \Security definitions\ theory Semi_Honest_Def imports CryptHOL.CryptHOL begin subsubsection \Security for deterministic functionalities\ locale sim_det_def = fixes R1 :: "'msg1 \ 'msg2 \ 'view1 spmf" and S1 :: "'msg1 \ 'out1 \ 'view1 spmf" and R2 :: "'msg1 \ 'msg2 \ 'view2 spmf" and S2 :: "'msg2 \ 'out2 \ 'view2 spmf" and funct :: "'msg1 \ 'msg2 \ ('out1 \ 'out2) spmf" and protocol :: "'msg1 \ 'msg2 \ ('out1 \ 'out2) spmf" assumes lossless_R1: "lossless_spmf (R1 m1 m2)" and lossless_S1: "lossless_spmf (S1 m1 out1)" and lossless_R2: "lossless_spmf (R2 m1 m2)" and lossless_S2: "lossless_spmf (S2 m2 out2)" and lossless_funct: "lossless_spmf (funct m1 m2)" begin type_synonym 'view' adversary_det = "'view' \ bool spmf" definition "correctness m1 m2 \ (protocol m1 m2 = funct m1 m2)" definition adv_P1 :: "'msg1 \ 'msg2 \ 'view1 adversary_det \ real" where "adv_P1 m1 m2 D \ \(spmf (R1 m1 m2 \ D) True) - spmf (funct m1 m2 \ (\ (o1, o2). S1 m1 o1 \ D)) True\" -definition "inf_theoretic_P1 m1 m2 \ (R1 m1 m2 = funct m1 m2 \ (\ (s1, s2). S1 m1 s1))" +definition "perfect_sec_P1 m1 m2 \ (R1 m1 m2 = funct m1 m2 \ (\ (s1, s2). S1 m1 s1))" definition adv_P2 :: "'msg1 \ 'msg2 \ 'view2 adversary_det \ real" where "adv_P2 m1 m2 D = \spmf (R2 m1 m2 \ (\ view. D view)) True - spmf (funct m1 m2 \ (\ (o1, o2). S2 m2 o2 \ (\ view. D view))) True\" -definition "inf_theoretic_P2 m1 m2 \ (R2 m1 m2 = funct m1 m2 \ (\ (s1, s2). S2 m2 s2))" +definition "perfect_sec_P2 m1 m2 \ (R2 m1 m2 = funct m1 m2 \ (\ (s1, s2). S2 m2 s2))" text \We also define the security games (for Party 1 and 2) used in EasyCrypt to define semi honest security for Party 1. We then show the two definitions are equivalent.\ definition P1_game_alt :: "'msg1 \ 'msg2 \ 'view1 adversary_det \ bool spmf" where "P1_game_alt m1 m2 D = do { b \ coin_spmf; (out1, out2) \ funct m1 m2; rview :: 'view1 \ R1 m1 m2; sview :: 'view1 \ S1 m1 out1; b' \ D (if b then rview else sview); return_spmf (b = b')}" definition adv_P1_game :: "'msg1 \ 'msg2 \ 'view1 adversary_det \ real" where "adv_P1_game m1 m2 D = \2*(spmf (P1_game_alt m1 m2 D ) True) - 1\" text \We show the two definitions are equivalent\ lemma equiv_defs_P1: assumes lossless_D: "\ view. lossless_spmf ((D:: 'view1 adversary_det) view)" shows "adv_P1_game m1 m2 D = adv_P1 m1 m2 D" including monad_normalisation proof- have return_True_not_False: "spmf (return_spmf (b)) True = spmf (return_spmf (\ b)) False" for b by(cases b; auto) have lossless_ideal: "lossless_spmf ((funct m1 m2 \ (\(out1, out2). S1 m1 out1 \ (\sview. D sview \ (\b'. return_spmf (False = b'))))))" by(simp add: lossless_S1 lossless_funct lossless_weight_spmfD split_def lossless_D) have return: "spmf (funct m1 m2 \ (\(o1, o2). S1 m1 o1 \ D)) True = spmf (funct m1 m2 \ (\(o1, o2). S1 m1 o1 \ (\ view. D view \ (\ b. return_spmf b)))) True" by simp have "2*(spmf (P1_game_alt m1 m2 D ) True) - 1 = (spmf (R1 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S1 m1 out1 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)" by(simp add: spmf_bind integral_spmf_of_set adv_P1_game_def P1_game_alt_def spmf_of_set UNIV_bool bind_spmf_const lossless_R1 lossless_S1 lossless_funct lossless_weight_spmfD) hence "adv_P1_game m1 m2 D = \(spmf (R1 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S1 m1 out1 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)\" using adv_P1_game_def by simp also have "\(spmf (R1 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S1 m1 out1 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)\ = adv_P1 m1 m2 D" apply(simp only: adv_P1_def spmf_False_conv_True[symmetric] lossless_ideal; simp) by(simp only: return)(simp only: split_def spmf_bind return_True_not_False) ultimately show ?thesis by simp qed definition P2_game_alt :: "'msg1 \ 'msg2 \ 'view2 adversary_det \ bool spmf" where "P2_game_alt m1 m2 D = do { b \ coin_spmf; (out1, out2) \ funct m1 m2; rview :: 'view2 \ R2 m1 m2; sview :: 'view2 \ S2 m2 out2; b' \ D (if b then rview else sview); return_spmf (b = b')}" definition adv_P2_game :: "'msg1 \ 'msg2 \ 'view2 adversary_det \ real" where "adv_P2_game m1 m2 D = \2*(spmf (P2_game_alt m1 m2 D ) True) - 1\" lemma equiv_defs_P2: assumes lossless_D: "\ view. lossless_spmf ((D:: 'view2 adversary_det) view)" shows "adv_P2_game m1 m2 D = adv_P2 m1 m2 D" including monad_normalisation proof- have return_True_not_False: "spmf (return_spmf (b)) True = spmf (return_spmf (\ b)) False" for b by(cases b; auto) have lossless_ideal: "lossless_spmf ((funct m1 m2 \ (\(out1, out2). S2 m2 out2 \ (\sview. D sview \ (\b'. return_spmf (False = b'))))))" by(simp add: lossless_S2 lossless_funct lossless_weight_spmfD split_def lossless_D) have return: "spmf (funct m1 m2 \ (\(o1, o2). S2 m2 o2 \ D)) True = spmf (funct m1 m2 \ (\(o1, o2). S2 m2 o2 \ (\ view. D view \ (\ b. return_spmf b)))) True" by simp have "2*(spmf (P2_game_alt m1 m2 D ) True) - 1 = (spmf (R2 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S2 m2 out2 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)" by(simp add: spmf_bind integral_spmf_of_set adv_P1_game_def P2_game_alt_def spmf_of_set UNIV_bool bind_spmf_const lossless_R2 lossless_S2 lossless_funct lossless_weight_spmfD) hence "adv_P2_game m1 m2 D = \(spmf (R2 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S2 m2 out2 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)\" using adv_P2_game_def by simp also have "\(spmf (R2 m1 m2 \ (\rview. D rview \ (\(b':: bool). return_spmf (True = b'))))) True - (1 - (spmf (funct m1 m2 \ (\(out1, out2). S2 m2 out2 \ (\sview. D sview \ (\b'. return_spmf (False = b')))))) True)\ = adv_P2 m1 m2 D" apply(simp only: adv_P2_def spmf_False_conv_True[symmetric] lossless_ideal; simp) by(simp only: return)(simp only: split_def spmf_bind return_True_not_False) ultimately show ?thesis by simp qed end subsubsection \ Security definitions for non deterministic functionalities \ locale sim_non_det_def = fixes R1 :: "'msg1 \ 'msg2 \ ('view1 \ ('out1 \ 'out2)) spmf" and S1 :: "'msg1 \ 'out1 \ 'view1 spmf" and Out1 :: "'msg1 \ 'msg2 \ 'out1 \ ('out1 \ 'out2) spmf" \ \takes the input of the other party so can form the outputs of parties\ and R2 :: "'msg1 \ 'msg2 \ ('view2 \ ('out1 \ 'out2)) spmf" and S2 :: "'msg2 \ 'out2 \ 'view2 spmf" and Out2 :: "'msg2 \ 'msg1 \ 'out2 \ ('out1 \ 'out2) spmf" and funct :: "'msg1 \ 'msg2 \ ('out1 \ 'out2) spmf" begin type_synonym ('view', 'out1', 'out2') adversary_non_det = "('view' \ ('out1' \ 'out2')) \ bool spmf" definition Ideal1 :: "'msg1 \ 'msg2 \ 'out1 \ ('view1 \ ('out1 \ 'out2)) spmf" where "Ideal1 m1 m2 out1 = do { view1 :: 'view1 \ S1 m1 out1; out1 \ Out1 m1 m2 out1; return_spmf (view1, out1)}" definition Ideal2 :: "'msg2 \ 'msg1 \ 'out2 \ ('view2 \ ('out1 \ 'out2)) spmf" where "Ideal2 m2 m1 out2 = do { view2 :: 'view2 \ S2 m2 out2; out2 \ Out2 m2 m1 out2; return_spmf (view2, out2)}" definition adv_P1 :: "'msg1 \ 'msg2 \ ('view1, 'out1, 'out2) adversary_non_det \ real" where "adv_P1 m1 m2 D \ \(spmf (R1 m1 m2 \ (\ view. D view)) True) - spmf (funct m1 m2 \ (\ (o1, o2). Ideal1 m1 m2 o1 \ (\ view. D view))) True\" -definition "inf_theoretic_P1 m1 m2 \ (R1 m1 m2 = funct m1 m2 \ (\ (s1, s2). Ideal1 m1 m2 s1))" +definition "perfect_sec_P1 m1 m2 \ (R1 m1 m2 = funct m1 m2 \ (\ (s1, s2). Ideal1 m1 m2 s1))" definition adv_P2 :: "'msg1 \ 'msg2 \ ('view2, 'out1, 'out2) adversary_non_det \ real" where "adv_P2 m1 m2 D = \spmf (R2 m1 m2 \ (\ view. D view)) True - spmf (funct m1 m2 \ (\ (o1, o2). Ideal2 m2 m1 o2 \ (\ view. D view))) True\" -definition "inf_theoretic_P2 m1 m2 \ (R2 m1 m2 = funct m1 m2 \ (\ (s1, s2). Ideal2 m2 m1 s2))" +definition "perfect_sec_P2 m1 m2 \ (R2 m1 m2 = funct m1 m2 \ (\ (s1, s2). Ideal2 m2 m1 s2))" end subsubsection \ Secret sharing schemes \ locale secret_sharing_scheme = fixes share :: "'input_out \ ('share \ 'share) spmf" and reconstruct :: "('share \ 'share) \ 'input_out spmf" and F :: "('input_out \ 'input_out \ 'input_out spmf) set" begin definition "sharing_correct input \ (share input \ (\ (s1,s2). reconstruct (s1,s2)) = return_spmf input)" definition "correct_share_eval input1 input2 \ (\ gate_eval \ F. \ gate_protocol :: ('share \ 'share) \ ('share \ 'share) \ ('share \ 'share) spmf. share input1 \ (\ (s1,s2). share input2 \ (\ (s3,s4). gate_protocol (s1,s3) (s2,s4) \ (\ (S1,S2). reconstruct (S1,S2)))) = gate_eval input1 input2)" end end