diff --git a/thys/Constructive_Cryptography_CM/Asymptotic_Security.thy b/thys/Constructive_Cryptography_CM/Asymptotic_Security.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Asymptotic_Security.thy @@ -0,0 +1,181 @@ +theory Asymptotic_Security imports Concrete_Security begin + +section \Asymptotic security definition\ + +locale constructive_security_obsf' = + fixes real_resource :: "security \ ('a + 'e, 'b + 'f) resource" + and ideal_resource :: "security \ ('c + 'e, 'd + 'f) resource" + and sim :: "security \ ('a, 'b, 'c, 'd) converter" + and \_real :: "security \ ('a, 'b) \" + and \_ideal :: "security \ ('c, 'd) \" + and \_common :: "security \ ('e, 'f) \" + and \ :: "security \ ('a + 'e, 'b + 'f) distinguisher_obsf" + assumes constructive_security_aux_obsf: "\\. + constructive_security_aux_obsf (real_resource \) (ideal_resource \) (sim \) (\_real \) (\_ideal \) (\_common \) 0" + and adv: "\ \\. exception_\ (\_real \ \\<^sub>\ \_common \) \g \ \ \ \ + \ negligible (\\. advantage (\ \) (obsf_resource (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \)) (obsf_resource (real_resource \)))" +begin + +sublocale constructive_security_aux_obsf + "real_resource \" + "ideal_resource \" + "sim \" + "\_real \" + "\_ideal \" + "\_common \" + "0" + for \ by(rule constructive_security_aux_obsf) + +lemma constructive_security_obsf'D: + "constructive_security_obsf (real_resource \) (ideal_resource \) (sim \) (\_real \) (\_ideal \) (\_common \) (\ \) + (advantage (\ \) (obsf_resource (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \)) (obsf_resource (real_resource \)))" + by(rule constructive_security_obsf_refl) + +end + +lemma constructive_security_obsf'I: + assumes "\\. constructive_security_obsf (real_resource \) (ideal_resource \) (sim \) (\_real \) (\_ideal \) (\_common \) (\ \) (adv \)" + and "(\\. exception_\ (\_real \ \\<^sub>\ \_common \) \g \ \ \) \ negligible adv" + shows "constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common \" +proof - + interpret constructive_security_obsf + "real_resource \" + "ideal_resource \" + "sim \" + "\_real \" + "\_ideal \" + "\_common \" + "\ \" + "adv \" + for \ by fact + show ?thesis + proof + show "negligible (\\. advantage (\ \) (obsf_resource (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \)) (obsf_resource (real_resource \)))" + if "\\. exception_\ (\_real \ \\<^sub>\ \_common \) \g \ \ \" using assms(2)[OF that] + by(rule negligible_mono)(auto intro!: eventuallyI landau_o.big_mono simp add: advantage_nonneg adv_nonneg adv[OF that]) + qed(rule WT_intro pfinite_intro order_refl)+ +qed + +lemma constructive_security_obsf'_into_constructive_security: + assumes "\\ :: security \ ('a + 'b, 'c + 'd) distinguisher_obsf. + \ \\. interaction_bounded_by (\_. True) (\ \) (bound \); + \\. lossless \ plossless_gpv (exception_\ (\_real \ \\<^sub>\ \_common \)) (\ \) \ + \ constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common \" + and correct: "\cnv. \\. (\\. \_ideal \ \\<^sub>\ \_common \ \g \ \ \) \ + (\\. interaction_any_bounded_by (\ \) (bound \)) \ + (\\. lossless \ plossless_gpv (\_ideal \ \\<^sub>\ \_common \) (\ \)) \ + (\\. wiring (\_ideal \) (\_real \) (cnv \) (w \)) \ + Negligible.negligible (\\. advantage (\ \) (ideal_resource \) (cnv \ |\<^sub>= 1\<^sub>C \ real_resource \))" + shows "constructive_security real_resource ideal_resource sim \_real \_ideal \_common bound lossless w" +proof + interpret constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common \\_. Done undefined\ + by(rule assms) simp_all + show "\_real \ \\<^sub>\ \_common \ \res real_resource \ \" + and "\_ideal \ \\<^sub>\ \_common \ \res ideal_resource \ \" + and "\_real \, \_ideal \ \\<^sub>C sim \ \" for \ by(rule WT_intro)+ + + show "\cnv. \\. (\\. \_ideal \ \\<^sub>\ \_common \ \g \ \ \) \ + (\\. interaction_any_bounded_by (\ \) (bound \)) \ + (\\. lossless \ plossless_gpv (\_ideal \ \\<^sub>\ \_common \) (\ \)) \ + (\\. wiring (\_ideal \) (\_real \) (cnv \) (w \)) \ + Negligible.negligible (\\. advantage (\ \) (ideal_resource \) (cnv \ |\<^sub>= 1\<^sub>C \ real_resource \))" + by fact +next + fix \ :: "security \ ('a + 'b, 'c + 'd) distinguisher" + assume WT_adv [WT_intro]: "\\. \_real \ \\<^sub>\ \_common \ \g \ \ \" + and bound [interaction_bound]: "\\. interaction_any_bounded_by (\ \) (bound \)" + and lossless: "\\. lossless \ plossless_gpv (\_real \ \\<^sub>\ \_common \) (\ \)" + let ?\ = "\\. obsf_distinguisher (\ \)" + interpret constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common ?\ + proof(rule assms) + show "interaction_any_bounded_by (?\ \) (bound \)" for \ by(rule interaction_bound)+ + show "plossless_gpv (exception_\ (\_real \ \\<^sub>\ \_common \)) (?\ \)" if lossless for \ + using WT_adv[of \] lossless that by(simp) + qed + have "negligible (\\. advantage (?\ \) (obsf_resource (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \)) (obsf_resource (real_resource \)))" + by(rule adv)(rule WT_intro)+ + then show "negligible (\\. advantage (\ \) (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \) (real_resource \))" + unfolding advantage_obsf_distinguisher . +qed + +subsection \Composition theorems\ + +theorem constructive_security_obsf'_composability: + fixes real + assumes "constructive_security_obsf' middle ideal sim_inner \_middle \_inner \_common (\\. absorb (\ \) (obsf_converter (sim_outer \ |\<^sub>= 1\<^sub>C)))" + assumes "constructive_security_obsf' real middle sim_outer \_real \_middle \_common \" + shows "constructive_security_obsf' real ideal (\\. sim_outer \ \ sim_inner \) \_real \_inner \_common \" +proof(rule constructive_security_obsf'I) + let ?\ = "\\. absorb (\ \) (obsf_converter (sim_outer \ |\<^sub>= 1\<^sub>C))" + interpret inner: constructive_security_obsf' middle ideal sim_inner \_middle \_inner \_common ?\ by fact + interpret outer: constructive_security_obsf' real middle sim_outer \_real \_middle \_common \ by fact + + let ?adv1 = "\\. advantage (?\ \) (obsf_resource (sim_inner \ |\<^sub>= 1\<^sub>C \ ideal \)) (obsf_resource (middle \))" + let ?adv2 = "\\. advantage (\ \) (obsf_resource (sim_outer \ |\<^sub>= 1\<^sub>C \ middle \)) (obsf_resource (real \))" + let ?adv = "\\. ?adv1 \ + ?adv2 \" + show "constructive_security_obsf (real \) (ideal \) (sim_outer \ \ sim_inner \) (\_real \) (\_inner \) (\_common \) (\ \) (?adv \)" for \ + using inner.constructive_security_obsf'D outer.constructive_security_obsf'D + by(rule constructive_security_obsf_composability) + assume [WT_intro]: "exception_\ (\_real \ \\<^sub>\ \_common \) \g \ \ \" for \ + have "negligible ?adv1" by(rule inner.adv)(rule WT_intro)+ + also have "negligible ?adv2" by(rule outer.adv)(rule WT_intro)+ + finally (negligible_plus) show "negligible ?adv" . +qed + +theorem constructive_security_obsf'_lifting: (* TODO: generalize! *) + assumes sec: "constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common (\\. absorb (\ \) (obsf_converter (1\<^sub>C |\<^sub>= conv \)))" + assumes WT_conv [WT_intro]: "\\. \_common' \, \_common \ \\<^sub>C conv \ \" + and pfinite [pfinite_intro]: "\\. pfinite_converter (\_common' \) (\_common \) (conv \)" + shows "constructive_security_obsf' + (\\. 1\<^sub>C |\<^sub>= conv \ \ real_resource \) (\\. 1\<^sub>C |\<^sub>= conv \ \ ideal_resource \) sim + \_real \_ideal \_common' \" +proof(rule constructive_security_obsf'I) + let ?\ = "\\. absorb (\ \) (obsf_converter (1\<^sub>C |\<^sub>= conv \))" + interpret constructive_security_obsf' real_resource ideal_resource sim \_real \_ideal \_common ?\ by fact + let ?adv = "\\. advantage (?\ \) (obsf_resource (sim \ |\<^sub>= 1\<^sub>C \ ideal_resource \)) (obsf_resource (real_resource \))" + + fix \ :: security + show "constructive_security_obsf (1\<^sub>C |\<^sub>= conv \ \ real_resource \) (1\<^sub>C |\<^sub>= conv \ \ ideal_resource \) (sim \) + (\_real \) (\_ideal \) (\_common' \) (\ \) + (?adv \)" + using constructive_security_obsf.constructive_security_aux_obsf[OF constructive_security_obsf'D] + constructive_security_obsf.constructive_security_sim_obsf[OF constructive_security_obsf'D] + by(rule constructive_security_obsf_lifting_usr)(rule WT_intro pfinite_intro)+ + show "negligible ?adv" if [WT_intro]: "\\. exception_\ (\_real \ \\<^sub>\ \_common' \) \g \ \ \" + by(rule adv)(rule WT_intro)+ +qed + +theorem constructive_security_obsf'_trivial: + fixes res + assumes [WT_intro]: "\\. \ \ \\<^sub>\ \_common \ \res res \ \" + shows "constructive_security_obsf' res res (\_. 1\<^sub>C) \ \ \_common \" +proof(rule constructive_security_obsf'I) + show "constructive_security_obsf (res \) (res \) 1\<^sub>C (\ \) (\ \) (\_common \) (\ \) 0" for \ + using assms by(rule constructive_security_obsf_trivial) +qed simp + +theorem parallel_constructive_security_obsf': + assumes "constructive_security_obsf' real1 ideal1 sim1 \_real1 \_inner1 \_common1 (\\. absorb (\ \) (obsf_converter (parallel_wiring \ parallel_converter 1\<^sub>C (converter_of_resource (sim2 \ |\<^sub>= 1\<^sub>C \ ideal2 \)))))" + (is "constructive_security_obsf' _ _ _ _ _ _ ?\1") + assumes "constructive_security_obsf' real2 ideal2 sim2 \_real2 \_inner2 \_common2 (\\. absorb (\ \) (obsf_converter (parallel_wiring \ parallel_converter (converter_of_resource (real1 \)) 1\<^sub>C)))" + (is "constructive_security_obsf' _ _ _ _ _ _ ?\2") + shows "constructive_security_obsf' (\\. parallel_wiring \ real1 \ \ real2 \) (\\. parallel_wiring \ ideal1 \ \ ideal2 \) (\\. sim1 \ |\<^sub>= sim2 \) + (\\. \_real1 \ \\<^sub>\ \_real2 \) (\\. \_inner1 \ \\<^sub>\ \_inner2 \) (\\. \_common1 \ \\<^sub>\ \_common2 \) \" +proof(rule constructive_security_obsf'I) + interpret sec1: constructive_security_obsf' real1 ideal1 sim1 \_real1 \_inner1 \_common1 ?\1 by fact + interpret sec2: constructive_security_obsf' real2 ideal2 sim2 \_real2 \_inner2 \_common2 ?\2 by fact + let ?adv1 = "\\. advantage (?\1 \) (obsf_resource (sim1 \ |\<^sub>= 1\<^sub>C \ ideal1 \)) (obsf_resource (real1 \))" + let ?adv2 = "\\. advantage (?\2 \) (obsf_resource (sim2 \ |\<^sub>= 1\<^sub>C \ ideal2 \)) (obsf_resource (real2 \))" + let ?adv = "\\. ?adv1 \ + ?adv2 \" + show "constructive_security_obsf (parallel_wiring \ real1 \ \ real2 \) (parallel_wiring \ ideal1 \ \ ideal2 \) + (sim1 \ |\<^sub>= sim2 \) (\_real1 \ \\<^sub>\ \_real2 \) (\_inner1 \ \\<^sub>\ \_inner2 \) (\_common1 \ \\<^sub>\ \_common2 \) (\ \) + (?adv \)" for \ + using sec1.constructive_security_obsf'D sec2.constructive_security_obsf'D + by(rule parallel_constructive_security_obsf) + assume [WT_intro]: "exception_\ ((\_real1 \ \\<^sub>\ \_real2 \) \\<^sub>\ (\_common1 \ \\<^sub>\ \_common2 \)) \g \ \ \" for \ + have "negligible ?adv1" by(rule sec1.adv)(rule WT_intro)+ + also have "negligible ?adv2" by(rule sec2.adv)(rule WT_intro)+ + finally (negligible_plus) show "negligible ?adv" . +qed + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Concrete_Security.thy b/thys/Constructive_Cryptography_CM/Concrete_Security.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Concrete_Security.thy @@ -0,0 +1,482 @@ +theory Concrete_Security +imports + Observe_Failure + Construction_Utility +begin + +section \Concrete security definition\ + +locale constructive_security_aux_obsf = + fixes real_resource :: "('a + 'e, 'b + 'f) resource" + and ideal_resource :: "('c + 'e, 'd + 'f) resource" + and sim :: "('a, 'b, 'c, 'd) converter" + and \_real :: "('a, 'b) \" + and \_ideal :: "('c, 'd) \" + and \_common :: "('e, 'f) \" + and adv :: real + assumes WT_real [WT_intro]: "\_real \\<^sub>\ \_common \res real_resource \" + and WT_ideal [WT_intro]: "\_ideal \\<^sub>\ \_common \res ideal_resource \" + and WT_sim [WT_intro]: "\_real, \_ideal \\<^sub>C sim \" + and pfinite_sim [pfinite_intro]: "pfinite_converter \_real \_ideal sim" + and adv_nonneg: "0 \ adv" + +locale constructive_security_sim_obsf = + fixes real_resource :: "('a + 'e, 'b + 'f) resource" + and ideal_resource :: "('c + 'e, 'd + 'f) resource" + and sim :: "('a, 'b, 'c, 'd) converter" + and \_real :: "('a, 'b) \" + and \_common :: "('e, 'f) \" + and \ :: "('a + 'e, 'b + 'f) distinguisher_obsf" + and adv :: real + assumes adv: "\ exception_\ (\_real \\<^sub>\ \_common) \g \ \ \ + \ advantage \ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource (real_resource)) \ adv" + +locale constructive_security_obsf = constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal \_common adv + + constructive_security_sim_obsf real_resource ideal_resource sim \_real \_common \ adv + for real_resource :: "('a + 'e, 'b + 'f) resource" + and ideal_resource :: "('c + 'e, 'd + 'f) resource" + and sim :: "('a, 'b, 'c, 'd) converter" + and \_real :: "('a, 'b) \" + and \_ideal :: "('c, 'd) \" + and \_common :: "('e, 'f) \" + and \ :: "('a + 'e, 'b + 'f) distinguisher_obsf" + and adv :: real +begin + +lemma constructive_security_aux_obsf: "constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal \_common adv" .. +lemma constructive_security_sim_obsf: "constructive_security_sim_obsf real_resource ideal_resource sim \_real \_common \ adv" .. + +end + +context constructive_security_aux_obsf begin + +lemma constructive_security_obsf_refl: + "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common \ + (advantage \ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource (real_resource)))" + by unfold_locales(simp_all add: advantage_def WT_intro pfinite_intro) + +end + +lemma constructive_security_obsf_absorb_cong: + assumes sec: "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common (absorb \ cnv) adv" + and [WT_intro]: "exception_\ \, exception_\ (\_real \\<^sub>\ \_common) \\<^sub>C cnv \" "exception_\ \, exception_\ (\_real \\<^sub>\ \_common) \\<^sub>C cnv' \" "exception_\ \ \g \ \" + and cong: "exception_\ \, exception_\ (\_real \\<^sub>\ \_common) \\<^sub>C cnv \ cnv'" + shows "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common (absorb \ cnv') adv" +proof - + interpret constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common "absorb \ cnv" adv by fact + show ?thesis + proof + have "connect_obsf \ (cnv' \ obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) = connect_obsf \ (cnv \ obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource))" + "connect_obsf \ (cnv' \ obsf_resource real_resource) = connect_obsf \ (cnv \ obsf_resource real_resource)" + by(rule connect_eq_resource_cong eq_\_attach_on' WT_intro cong[symmetric] order_refl)+ + then have "advantage (absorb \ cnv') (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource) = + advantage (absorb \ cnv) (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource)" + unfolding advantage_def distinguish_attach[symmetric] by simp + also have "\ \ adv" by(rule adv)(rule WT_intro)+ + finally show "advantage (absorb \ cnv') (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource) \ adv" . + qed +qed + +lemma constructive_security_obsf_sim_cong: + assumes sec: "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common \ adv" + and cong: "\_real, \_ideal \\<^sub>C sim \ sim'" + and pfinite [pfinite_intro]: "pfinite_converter \_real \_ideal sim'" (* This could probably be derived from cong and sec.pfinite_sim *) + shows "constructive_security_obsf real_resource ideal_resource sim' \_real \_ideal \_common \ adv" +proof + interpret constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common \ adv by fact + show "\_real \\<^sub>\ \_common \res real_resource \" "\_ideal \\<^sub>\ \_common \res ideal_resource \" by(rule WT_intro)+ + from cong show [WT_intro]: "\_real, \_ideal \\<^sub>C sim' \" by(rule eq_\_converterD_WT1)(rule WT_intro) + show "pfinite_converter \_real \_ideal sim'" by fact + show "0 \ adv" by(rule adv_nonneg) + + assume WT [WT_intro]: "exception_\ (\_real \\<^sub>\ \_common) \g \ \" + have "connect_obsf \ (obsf_resource (sim' |\<^sub>= 1\<^sub>C \ ideal_resource)) = connect_obsf \ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource))" + by(rule connect_eq_resource_cong WT_intro obsf_resource_eq_\_cong eq_\_attach_on' parallel_converter2_eq_\_cong cong[symmetric] eq_\_converter_reflI | simp)+ + with adv[OF WT] + show "advantage \ (obsf_resource (sim' |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource) \ adv" + unfolding advantage_def by simp +qed + +lemma constructive_security_obsfI_core_rest [locale_witness]: + assumes "constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal (\_common_core \\<^sub>\ \_common_rest) adv" + and adv: "\ exception_\ (\_real \\<^sub>\ (\_common_core \\<^sub>\ \_common_rest)) \g \ \ \ + \ advantage \ (obsf_resource (sim |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ ideal_resource)) (obsf_resource (real_resource)) \ adv" + shows "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal (\_common_core \\<^sub>\ \_common_rest) \ adv" +proof - + interpret constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal "\_common_core \\<^sub>\ \_common_rest" by fact + show ?thesis + proof + assume A [WT_intro]: "exception_\ (\_real \\<^sub>\ (\_common_core \\<^sub>\ \_common_rest)) \g \ \" + hence outs: "outs_gpv (exception_\ (\_real \\<^sub>\ (\_common_core \\<^sub>\ \_common_rest))) \ \ outs_\ (\_real \\<^sub>\ (\_common_core \\<^sub>\ \_common_rest))" + unfolding WT_gpv_iff_outs_gpv by simp + have "connect_obsf \ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) = connect_obsf \ (obsf_resource (sim |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C \ ideal_resource))" + by(rule connect_cong_trace trace_eq_obsf_resourceI eq_resource_on_imp_trace_eq eq_\_attach_on')+ + (rule WT_intro parallel_converter2_eq_\_cong eq_\_converter_reflI parallel_converter2_id_id[symmetric] order_refl outs)+ + then show "advantage \ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource) \ adv" + using adv[OF A] by(simp add: advantage_def) + qed +qed + +subsection \Composition theorems\ + +theorem constructive_security_obsf_composability: + fixes real + assumes "constructive_security_obsf middle ideal sim_inner \_middle \_inner \_common (absorb \ (obsf_converter (sim_outer |\<^sub>= 1\<^sub>C))) adv1" + assumes "constructive_security_obsf real middle sim_outer \_real \_middle \_common \ adv2" + shows "constructive_security_obsf real ideal (sim_outer \ sim_inner) \_real \_inner \_common \ (adv1 + adv2)" +proof + let ?\ = "absorb \ (obsf_converter (sim_outer |\<^sub>= 1\<^sub>C))" + interpret inner: constructive_security_obsf middle ideal sim_inner \_middle \_inner \_common ?\ adv1 by fact + interpret outer: constructive_security_obsf real middle sim_outer \_real \_middle \_common \ adv2 by fact + + show "\_real \\<^sub>\ \_common \res real \" + and "\_inner \\<^sub>\ \_common \res ideal \" + and "\_real, \_inner \\<^sub>C sim_outer \ sim_inner \" by(rule WT_intro)+ + show "pfinite_converter \_real \_inner (sim_outer \ sim_inner)" by(rule pfinite_intro WT_intro)+ + show "0 \ adv1 + adv2" using inner.adv_nonneg outer.adv_nonneg by simp + + assume WT_adv[WT_intro]: "exception_\ (\_real \\<^sub>\ \_common) \g \ \" + have eq1: "connect_obsf (absorb \ (obsf_converter (sim_outer |\<^sub>= 1\<^sub>C))) (obsf_resource (sim_inner |\<^sub>= 1\<^sub>C \ ideal)) = + connect_obsf \ (obsf_resource (sim_outer \ sim_inner |\<^sub>= 1\<^sub>C \ ideal))" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\ add: parallel_converter2_comp1_out attach_compose) + apply(rule obsf_attach) + apply(rule pfinite_intro WT_intro)+ + done + have eq2: "connect_obsf (absorb \ (obsf_converter (sim_outer |\<^sub>= 1\<^sub>C))) (obsf_resource middle) = + connect_obsf \ (obsf_resource (sim_outer |\<^sub>= 1\<^sub>C \ middle))" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\ add: parallel_converter2_comp1_out attach_compose) + apply(rule obsf_attach) + apply(rule pfinite_intro WT_intro)+ + done + + have "advantage ?\ (obsf_resource (sim_inner |\<^sub>= 1\<^sub>C \ ideal)) (obsf_resource middle) \ adv1" + by(rule inner.adv)(rule WT_intro)+ + moreover have "advantage \ (obsf_resource (sim_outer |\<^sub>= 1\<^sub>C \ middle)) (obsf_resource real) \ adv2" + by(rule outer.adv)(rule WT_intro)+ + ultimately + show "advantage \ (obsf_resource (sim_outer \ sim_inner |\<^sub>= 1\<^sub>C \ ideal)) (obsf_resource real) \ adv1 + adv2" + by(auto simp add: advantage_def eq1 eq2 abs_diff_triangle_ineq2) +qed + +theorem constructive_security_obsf_lifting: + assumes sec: "constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal \_common adv" + and sec2: "exception_\ (\_real' \\<^sub>\ \_common') \g \ \ + \ constructive_security_sim_obsf real_resource ideal_resource sim \_real \_common (absorb \ (obsf_converter (w_adv_real |\<^sub>= w_usr))) adv" + (is "_ \ constructive_security_sim_obsf _ _ _ _ _ ?\ _") + assumes WT_usr [WT_intro]: "\_common', \_common \\<^sub>C w_usr \" + and pfinite [pfinite_intro]: "pfinite_converter \_common' \_common w_usr" + and WT_adv_real [WT_intro]: "\_real', \_real \\<^sub>C w_adv_real \" + and WT_w_adv_ideal [WT_intro]: "\_ideal', \_ideal \\<^sub>C w_adv_ideal \" + and WT_adv_ideal_inv [WT_intro]: "\_ideal, \_ideal' \\<^sub>C w_adv_ideal_inv \" + and ideal_inverse: "\_ideal, \_ideal \\<^sub>C w_adv_ideal_inv \ w_adv_ideal \ 1\<^sub>C" + and pfinite_real [pfinite_intro]: "pfinite_converter \_real' \_real w_adv_real" + and pfinite_ideal [pfinite_intro]: "pfinite_converter \_ideal \_ideal' w_adv_ideal_inv" + shows "constructive_security_obsf (w_adv_real |\<^sub>= w_usr \ real_resource) (w_adv_ideal |\<^sub>= w_usr \ ideal_resource) (w_adv_real \ sim \ w_adv_ideal_inv) \_real' \_ideal' \_common' \ adv" + (is "constructive_security_obsf ?real ?ideal ?sim ?\_real ?\_ideal _ _ _") +proof + interpret constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal \_common by fact + show "\_real' \\<^sub>\ \_common' \res ?real \" + and "\_ideal' \\<^sub>\ \_common' \res ?ideal \" + and "\_real', \_ideal' \\<^sub>C ?sim \" by(rule WT_intro)+ + show "pfinite_converter \_real' \_ideal' ?sim" by(rule pfinite_intro WT_intro)+ + show "0 \ adv" by(rule adv_nonneg) + + assume WT_adv [WT_intro]: "exception_\ (\_real' \\<^sub>\ \_common') \g \ \" + then interpret constructive_security_sim_obsf real_resource ideal_resource sim \_real \_common ?\ adv by(rule sec2) + + have *: "advantage ?\ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) (obsf_resource real_resource) \ adv" + by(rule adv)(rule WT_intro)+ + + have ideal: "connect_obsf ?\ (obsf_resource (sim |\<^sub>= 1\<^sub>C \ ideal_resource)) = + connect_obsf \ (obsf_resource (?sim |\<^sub>= 1\<^sub>C \ ?ideal))" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule eq_resource_on_trans[OF obsf_attach]) + apply(rule pfinite_intro WT_intro)+ + apply(rule obsf_resource_eq_\_cong) + apply(fold attach_compose) + apply(unfold comp_converter_parallel2) + apply(rule eq_\_attach_on') + apply(rule WT_intro) + apply(rule parallel_converter2_eq_\_cong) + apply(unfold comp_converter_assoc) + apply(rule eq_\_comp_cong) + apply(rule eq_\_converter_reflI; rule WT_intro) + apply(rule eq_\_converter_trans[rotated]) + apply(rule eq_\_comp_cong) + apply(rule eq_\_converter_reflI; rule WT_intro) + apply(rule ideal_inverse[symmetric]) + apply(unfold comp_converter_id_right comp_converter_id_left) + apply(rule eq_\_converter_reflI; rule WT_intro)+ + apply simp + apply(rule WT_intro)+ + done + have real: "connect_obsf ?\ (obsf_resource real_resource) = connect_obsf \ (obsf_resource ?real)" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule obsf_attach) + apply(rule pfinite_intro WT_intro)+ + done + show "advantage \ (obsf_resource ((?sim |\<^sub>= 1\<^sub>C) \ ?ideal)) (obsf_resource ?real) \ adv" using * + unfolding advantage_def ideal[symmetric] real[symmetric] . +qed + +corollary constructive_security_obsf_lifting_: + assumes sec: "constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common (absorb \ (obsf_converter (w_adv_real |\<^sub>= w_usr))) adv" + assumes WT_usr [WT_intro]: "\_common', \_common \\<^sub>C w_usr \" + and pfinite [pfinite_intro]: "pfinite_converter \_common' \_common w_usr" + and WT_adv_real [WT_intro]: "\_real', \_real \\<^sub>C w_adv_real \" + and WT_w_adv_ideal [WT_intro]: "\_ideal', \_ideal \\<^sub>C w_adv_ideal \" + and WT_adv_ideal_inv [WT_intro]: "\_ideal, \_ideal' \\<^sub>C w_adv_ideal_inv \" + and ideal_inverse: "\_ideal, \_ideal \\<^sub>C w_adv_ideal_inv \ w_adv_ideal \ 1\<^sub>C" + and pfinite_real [pfinite_intro]: "pfinite_converter \_real' \_real w_adv_real" + and pfinite_ideal [pfinite_intro]: "pfinite_converter \_ideal \_ideal' w_adv_ideal_inv" + shows "constructive_security_obsf (w_adv_real |\<^sub>= w_usr \ real_resource) (w_adv_ideal |\<^sub>= w_usr \ ideal_resource) (w_adv_real \ sim \ w_adv_ideal_inv) \_real' \_ideal' \_common' \ adv" +proof - + interpret constructive_security_obsf real_resource ideal_resource sim \_real \_ideal \_common "absorb \ (obsf_converter (w_adv_real |\<^sub>= w_usr))" adv by fact + from constructive_security_aux_obsf constructive_security_sim_obsf assms(2-) + show ?thesis by(rule constructive_security_obsf_lifting) +qed + +theorem constructive_security_obsf_lifting_usr: + assumes sec: "constructive_security_aux_obsf real_resource ideal_resource sim \_real \_ideal \_common adv" + and sec2: "exception_\ (\_real \\<^sub>\ \_common') \g \ \ + \ constructive_security_sim_obsf real_resource ideal_resource sim \_real \_common (absorb \ (obsf_converter (1\<^sub>C |\<^sub>= conv))) adv" + and WT_conv [WT_intro]: "\_common', \_common \\<^sub>C conv \" + and pfinite [pfinite_intro]: "pfinite_converter \_common' \_common conv" + shows "constructive_security_obsf (1\<^sub>C |\<^sub>= conv \ real_resource) (1\<^sub>C |\<^sub>= conv \ ideal_resource) sim \_real \_ideal \_common' \ adv" + by(rule constructive_security_obsf_lifting[OF sec sec2, where ?w_adv_ideal="1\<^sub>C" and ?w_adv_ideal_inv="1\<^sub>C", simplified comp_converter_id_left comp_converter_id_right]) + (rule WT_intro pfinite_intro id_converter_eq_self order_refl | assumption)+ + +theorem constructive_security_obsf_lifting2: + assumes sec: "constructive_security_aux_obsf real_resource ideal_resource sim (\_real1 \\<^sub>\ \_real2) (\_ideal1 \\<^sub>\ \_ideal2) \_common adv" + and sec2: "exception_\ ((\_real1 \\<^sub>\ \_real2) \\<^sub>\ \_common') \g \ \ + \ constructive_security_sim_obsf real_resource ideal_resource sim (\_real1 \\<^sub>\ \_real2) \_common (absorb \ (obsf_converter ((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= conv))) adv" + assumes WT_conv [WT_intro]: "\_common', \_common \\<^sub>C conv \" + and pfinite [pfinite_intro]: "pfinite_converter \_common' \_common conv" + shows "constructive_security_obsf ((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= conv \ real_resource) ((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= conv \ ideal_resource) sim (\_real1 \\<^sub>\ \_real2) (\_ideal1 \\<^sub>\ \_ideal2) \_common' \ adv" + (is "constructive_security_obsf ?real ?ideal _ ?\_real ?\_ideal _ _ _") +proof - + interpret constructive_security_aux_obsf real_resource ideal_resource sim "\_real1 \\<^sub>\ \_real2" "\_ideal1 \\<^sub>\ \_ideal2" \_common adv by fact + have sim [unfolded comp_converter_id_left]: "\_real1 \\<^sub>\ \_real2,\_ideal1 \\<^sub>\ \_ideal2 \\<^sub>C (1\<^sub>C |\<^sub>= 1\<^sub>C) \ sim \ 1\<^sub>C \ sim" + by(rule eq_\_comp_cong)(rule parallel_converter2_id_id eq_\_converter_reflI WT_intro)+ + show ?thesis + apply(rule constructive_security_obsf_sim_cong) + apply(rule constructive_security_obsf_lifting[OF sec sec2, where ?w_adv_ideal="1\<^sub>C |\<^sub>= 1\<^sub>C" and ?w_adv_ideal_inv="1\<^sub>C", unfolded comp_converter_id_left comp_converter_id_right]) + apply(assumption|rule WT_intro sim pfinite_intro parallel_converter2_id_id)+ + done +qed + +theorem constructive_security_obsf_trivial: + fixes res + assumes [WT_intro]: "\ \\<^sub>\ \_common \res res \" + shows "constructive_security_obsf res res 1\<^sub>C \ \ \_common \ 0" +proof + show "\ \\<^sub>\ \_common \res res \" and "\, \ \\<^sub>C 1\<^sub>C \" by(rule WT_intro)+ + show "pfinite_converter \ \ 1\<^sub>C" by(rule pfinite_intro) + + assume WT [WT_intro]: "exception_\ (\ \\<^sub>\ \_common) \g \ \" + have "connect_obsf \ (obsf_resource (1\<^sub>C |\<^sub>= 1\<^sub>C \ res)) = connect_obsf \ (obsf_resource (1\<^sub>C \ res))" + by(rule connect_eq_resource_cong[OF WT])(fastforce intro: WT_intro eq_\_attach_on' obsf_resource_eq_\_cong parallel_converter2_id_id)+ + then show "advantage \ (obsf_resource (1\<^sub>C |\<^sub>= 1\<^sub>C \ res)) (obsf_resource res) \ 0" + unfolding advantage_def by simp +qed simp + +lemma parallel_constructive_security_aux_obsf [locale_witness]: + assumes "constructive_security_aux_obsf real1 ideal1 sim1 \_real1 \_inner1 \_common1 adv1" + assumes "constructive_security_aux_obsf real2 ideal2 sim2 \_real2 \_inner2 \_common2 adv2" + shows "constructive_security_aux_obsf (parallel_wiring \ real1 \ real2) (parallel_wiring \ ideal1 \ ideal2) (sim1 |\<^sub>= sim2) + (\_real1 \\<^sub>\ \_real2) (\_inner1 \\<^sub>\ \_inner2) (\_common1 \\<^sub>\ \_common2) + (adv1 + adv2)" +proof + interpret sec1: constructive_security_aux_obsf real1 ideal1 sim1 \_real1 \_inner1 \_common1 adv1 by fact + interpret sec2: constructive_security_aux_obsf real2 ideal2 sim2 \_real2 \_inner2 \_common2 adv2 by fact + + show "(\_real1 \\<^sub>\ \_real2) \\<^sub>\ (\_common1 \\<^sub>\ \_common2) \res parallel_wiring \ real1 \ real2 \" + and "(\_inner1 \\<^sub>\ \_inner2) \\<^sub>\ (\_common1 \\<^sub>\ \_common2) \res parallel_wiring \ ideal1 \ ideal2 \" + and "\_real1 \\<^sub>\ \_real2, \_inner1 \\<^sub>\ \_inner2 \\<^sub>C sim1 |\<^sub>= sim2 \" by(rule WT_intro)+ + show "pfinite_converter (\_real1 \\<^sub>\ \_real2) (\_inner1 \\<^sub>\ \_inner2) (sim1 |\<^sub>= sim2)" by(rule pfinite_intro)+ + show "0 \ adv1 + adv2" using sec1.adv_nonneg sec2.adv_nonneg by simp +qed + +theorem parallel_constructive_security_obsf: + assumes "constructive_security_obsf real1 ideal1 sim1 \_real1 \_inner1 \_common1 (absorb \ (obsf_converter (parallel_wiring \ parallel_converter 1\<^sub>C (converter_of_resource (sim2 |\<^sub>= 1\<^sub>C \ ideal2))))) adv1" + (is "constructive_security_obsf _ _ _ _ _ _ ?\1 _") + assumes "constructive_security_obsf real2 ideal2 sim2 \_real2 \_inner2 \_common2 (absorb \ (obsf_converter (parallel_wiring \ parallel_converter (converter_of_resource real1) 1\<^sub>C))) adv2" + (is "constructive_security_obsf _ _ _ _ _ _ ?\2 _") + shows "constructive_security_obsf (parallel_wiring \ real1 \ real2) (parallel_wiring \ ideal1 \ ideal2) (sim1 |\<^sub>= sim2) + (\_real1 \\<^sub>\ \_real2) (\_inner1 \\<^sub>\ \_inner2) (\_common1 \\<^sub>\ \_common2) + \ (adv1 + adv2)" +proof - + interpret sec1: constructive_security_obsf real1 ideal1 sim1 \_real1 \_inner1 \_common1 ?\1 adv1 by fact + interpret sec2: constructive_security_obsf real2 ideal2 sim2 \_real2 \_inner2 \_common2 ?\2 adv2 by fact + + show ?thesis + proof + assume WT [WT_intro]: "exception_\ ((\_real1 \\<^sub>\ \_real2) \\<^sub>\ (\_common1 \\<^sub>\ \_common2)) \g \ \" + + have **: "outs_\ ((\_real1 \\<^sub>\ \_real2) \\<^sub>\ (\_common1 \\<^sub>\ \_common2)) \\<^sub>R + ((1\<^sub>C |\<^sub>= sim2) |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ real1 \ ideal2 \ + parallel_wiring \ (converter_of_resource real1 |\<^sub>\ 1\<^sub>C) \ sim2 |\<^sub>= 1\<^sub>C \ ideal2" + unfolding comp_parallel_wiring + by(rule eq_resource_on_trans, rule eq_\_attach_on[where conv'="parallel_wiring \ (1\<^sub>C |\<^sub>= sim2 |\<^sub>= 1\<^sub>C)"] + , (rule WT_intro)+, rule eq_\_comp_cong, rule eq_\_converter_mono) + (auto simp add: le_\_def attach_compose attach_parallel2 attach_converter_of_resource_conv_parallel_resource + intro: WT_intro parallel_converter2_eq_\_cong parallel_converter2_id_id eq_\_converter_reflI) + + have ideal2: + "connect_obsf ?\2 (obsf_resource (sim2 |\<^sub>= 1\<^sub>C \ ideal2)) = + connect_obsf \ (obsf_resource ((1\<^sub>C |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ real1 \ ideal2))" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule eq_resource_on_trans[OF obsf_attach]) + apply(rule pfinite_intro WT_intro)+ + apply(rule obsf_resource_eq_\_cong) + apply(rule eq_resource_on_sym) + apply(subst attach_compose[symmetric]) + apply(rule **) + apply(rule WT_intro)+ + done + + have real2: "connect_obsf ?\2 (obsf_resource real2) = connect_obsf \ (obsf_resource (parallel_wiring \ real1 \ real2))" + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule eq_resource_on_trans[OF obsf_attach]) + apply(rule pfinite_intro WT_intro)+ + apply(rule obsf_resource_eq_\_cong) + apply(rule eq_resource_on_sym) + by(simp add: attach_compose attach_converter_of_resource_conv_parallel_resource)(rule WT_intro)+ + + have adv2: "advantage \ + (obsf_resource ((1\<^sub>C |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ real1 \ ideal2)) + (obsf_resource (parallel_wiring \ real1 \ real2)) \ adv2" + unfolding advantage_def ideal2[symmetric] real2[symmetric] by(rule sec2.adv[unfolded advantage_def])(rule WT_intro)+ + + have ideal1: + "connect_obsf ?\1 (obsf_resource (sim1 |\<^sub>= 1\<^sub>C \ ideal1)) = + connect_obsf \ (obsf_resource ((sim1 |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ ideal1 \ ideal2))" + proof - + have *:"((outs_\ \_real1 <+> outs_\ \_real2) <+> outs_\ \_common1 <+> outs_\ \_common2) \\<^sub>R + (sim1 |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ ideal1 \ ideal2 \ + parallel_wiring \ (1\<^sub>C |\<^sub>\ converter_of_resource (sim2 |\<^sub>= 1\<^sub>C \ ideal2)) \ sim1 |\<^sub>= 1\<^sub>C \ ideal1" + by(auto simp add: le_\_def comp_parallel_wiring' attach_compose attach_parallel2 attach_converter_of_resource_conv_parallel_resource2 intro: WT_intro) + show ?thesis + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule eq_resource_on_trans[OF obsf_attach]) + apply(rule pfinite_intro WT_intro)+ + apply(rule obsf_resource_eq_\_cong) + apply(rule eq_resource_on_sym) + by(simp add: *, (rule WT_intro)+) + qed + + have real1: "connect_obsf ?\1 (obsf_resource real1) = connect_obsf \ (obsf_resource ((1\<^sub>C |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ real1 \ ideal2))" + proof - + have *: "outs_\ ((\_real1 \\<^sub>\ \_real2) \\<^sub>\ (\_common1 \\<^sub>\ \_common2)) \\<^sub>R + parallel_wiring \ ((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= sim2 |\<^sub>= 1\<^sub>C) \ real1 \ ideal2 \ + parallel_wiring \ (1\<^sub>C |\<^sub>\ converter_of_resource (sim2 |\<^sub>= 1\<^sub>C \ ideal2 )) \ real1" + by(rule eq_resource_on_trans, rule eq_\_attach_on[where conv'="parallel_wiring \ (1\<^sub>C |\<^sub>= sim2 |\<^sub>= 1\<^sub>C)"] + , (rule WT_intro)+, rule eq_\_comp_cong, rule eq_\_converter_mono) + (auto simp add: le_\_def attach_compose attach_converter_of_resource_conv_parallel_resource2 attach_parallel2 + intro: WT_intro parallel_converter2_eq_\_cong parallel_converter2_id_id eq_\_converter_reflI) + + show ?thesis + unfolding distinguish_attach[symmetric] + apply(rule connect_eq_resource_cong) + apply(rule WT_intro) + apply(simp del: outs_plus_\) + apply(rule eq_resource_on_trans[OF obsf_attach]) + apply(rule pfinite_intro WT_intro)+ + apply(rule obsf_resource_eq_\_cong) + apply(rule eq_resource_on_sym) + apply(fold attach_compose) + apply(subst comp_parallel_wiring) + apply(rule *) + apply(rule WT_intro)+ + done + qed + + have adv1: "advantage \ + (obsf_resource ((sim1 |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ ideal1 \ ideal2)) + (obsf_resource ((1\<^sub>C |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ real1 \ ideal2)) \ adv1" + unfolding advantage_def ideal1[symmetric] real1[symmetric] by(rule sec1.adv[unfolded advantage_def])(rule WT_intro)+ + + from adv1 adv2 show "advantage \ (obsf_resource ((sim1 |\<^sub>= sim2) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ parallel_wiring \ ideal1 \ ideal2)) + (obsf_resource (parallel_wiring \ real1 \ real2)) \ adv1 + adv2" + by(auto simp add: advantage_def) + qed +qed + +theorem parallel_constructive_security_obsf_fuse: + assumes 1: "constructive_security_obsf real1 ideal1 sim1 (\_real1_core \\<^sub>\ \_real1_rest) (\_ideal1_core \\<^sub>\ \_ideal1_rest) (\_common1_core \\<^sub>\ \_common1_rest) (absorb \ (obsf_converter (fused_wiring \ parallel_converter 1\<^sub>C (converter_of_resource (sim2 |\<^sub>= 1\<^sub>C \ ideal2))))) adv1" + (is "constructive_security_obsf _ _ _ ?\_real1 ?\_ideal1 ?\_common1 ?\1 _") + assumes 2: "constructive_security_obsf real2 ideal2 sim2 (\_real2_core \\<^sub>\ \_real2_rest) (\_ideal2_core \\<^sub>\ \_ideal2_rest) (\_common2_core \\<^sub>\ \_common2_rest) (absorb \ (obsf_converter (fused_wiring \ parallel_converter (converter_of_resource real1) 1\<^sub>C))) adv2" + (is "constructive_security_obsf _ _ _ ?\_real2 ?\_ideal2 ?\_common2 ?\2 _") + shows "constructive_security_obsf (fused_wiring \ real1 \ real2) (fused_wiring \ ideal1 \ ideal2) + (parallel_wiring \ (sim1 |\<^sub>= sim2) \ parallel_wiring) + ((\_real1_core \\<^sub>\ \_real2_core) \\<^sub>\ (\_real1_rest \\<^sub>\ \_real2_rest)) + ((\_ideal1_core \\<^sub>\ \_ideal2_core) \\<^sub>\ (\_ideal1_rest \\<^sub>\ \_ideal2_rest)) + ((\_common1_core \\<^sub>\ \_common2_core) \\<^sub>\ (\_common1_rest \\<^sub>\ \_common2_rest)) + \ (adv1 + adv2)" +proof - + interpret sec1: constructive_security_obsf real1 ideal1 sim1 ?\_real1 ?\_ideal1 ?\_common1 ?\1 adv1 by fact + interpret sec2: constructive_security_obsf real2 ideal2 sim2 ?\_real2 ?\_ideal2 ?\_common2 ?\2 adv2 by fact + + have aux1: "constructive_security_aux_obsf real1 ideal1 sim1 ?\_real1 ?\_ideal1 ?\_common1 adv1" .. + have aux2: "constructive_security_aux_obsf real2 ideal2 sim2 ?\_real2 ?\_ideal2 ?\_common2 adv2" .. + + have sim: "constructive_security_sim_obsf (parallel_wiring \ real1 \ real2) (parallel_wiring \ ideal1 \ ideal2) (sim1 |\<^sub>= sim2) + (?\_real1 \\<^sub>\ ?\_real2) (?\_common1 \\<^sub>\ ?\_common2) + (absorb \ (obsf_converter (parallel_wiring |\<^sub>= parallel_wiring))) + (adv1 + adv2)" + if [WT_intro]: "exception_\ (((\_real1_core \\<^sub>\ \_real2_core) \\<^sub>\ (\_real1_rest \\<^sub>\ \_real2_rest)) \\<^sub>\ ((\_common1_core \\<^sub>\ \_common2_core) \\<^sub>\ (\_common1_rest \\<^sub>\ \_common2_rest))) \g \ \" + proof - + interpret constructive_security_obsf + "parallel_wiring \ real1 \ real2" + "parallel_wiring \ ideal1 \ ideal2" + "sim1 |\<^sub>= sim2" + "?\_real1 \\<^sub>\ ?\_real2" "?\_ideal1 \\<^sub>\ ?\_ideal2" "?\_common1 \\<^sub>\ ?\_common2" + "absorb \ (obsf_converter (parallel_wiring |\<^sub>= parallel_wiring))" + "adv1 + adv2" + apply(rule parallel_constructive_security_obsf) + apply(fold absorb_comp_converter) + apply(rule constructive_security_obsf_absorb_cong[OF 1]) + apply(rule WT_intro)+ + apply(unfold fused_wiring_def comp_converter_assoc) + apply(rule obsf_comp_converter) + apply(rule WT_intro pfinite_intro)+ + apply(rule constructive_security_obsf_absorb_cong[OF 2]) + apply(rule WT_intro)+ + apply(subst fused_wiring_def) + apply(unfold comp_converter_assoc) + apply(rule obsf_comp_converter) + apply(rule WT_intro pfinite_intro wiring_intro parallel_wiring_inverse)+ + done + show ?thesis .. + qed + show ?thesis + unfolding fused_wiring_def attach_compose + apply(rule constructive_security_obsf_lifting[where w_adv_ideal_inv=parallel_wiring]) + apply(rule parallel_constructive_security_aux_obsf[OF aux1 aux2]) + apply(erule sim) + apply(rule WT_intro pfinite_intro parallel_wiring_inverse)+ + done +qed + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Construction_Utility.thy b/thys/Constructive_Cryptography_CM/Construction_Utility.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Construction_Utility.thy @@ -0,0 +1,452 @@ +theory Construction_Utility + imports + Fused_Resource + State_Isomorphism +begin + +\ \Dummy converters that return a constant value on their external interface\ + +primcorec + ldummy_converter :: "('a \ 'b) \ ('i_cnv, 'o_cnv, 'i_res, 'o_res) converter \ + ('a + 'i_cnv, 'b + 'o_cnv, 'i_res, 'o_res) converter" + where + "run_converter (ldummy_converter f conv) = (\inp. case inp of + Inl x \ map_gpv (map_prod Inl (\_. ldummy_converter f conv)) id (Done (f x, ())) + | Inr x \ map_gpv (map_prod Inr (\c. ldummy_converter f c)) id (run_converter conv x))" + +primcorec + rdummy_converter :: "('a \ 'b) \ ('i_cnv, 'o_cnv, 'i_res, 'o_res) converter \ + ('i_cnv + 'a, 'o_cnv + 'b, 'i_res, 'o_res) converter" + where + "run_converter (rdummy_converter f conv) = (\inp. case inp of + Inl x \ map_gpv (map_prod Inl (\c. rdummy_converter f c)) id (run_converter conv x) + | Inr x \ map_gpv (map_prod Inr (\_. rdummy_converter f conv)) id (Done (f x, ())))" + +lemma ldummy_converter_of_callee: + "ldummy_converter f (converter_of_callee callee state) = + converter_of_callee (\s q. case_sum (\ql. Done (Inl (f ql), s)) (\qr. map_gpv (map_prod Inr id) id (callee s qr)) q) state" + apply (coinduction arbitrary: callee state) + apply(clarsimp intro!:rel_funI split!:sum.splits) + subgoal by blast + apply (simp add: gpv.rel_map map_prod_def split_def) + by (rule gpv.rel_mono_strong0[of "(=)" "(=)"]) (auto simp add: gpv.rel_eq) + +lemma rdummy_converter_of_callee: + "rdummy_converter f (converter_of_callee callee state) = + converter_of_callee (\s q. case_sum (\ql. map_gpv (map_prod Inl id) id (callee s ql)) (\qr. Done (Inr (f qr), s)) q) state" + apply (coinduction arbitrary: callee state) + apply(clarsimp intro!:rel_funI split!:sum.splits) + defer + subgoal by blast + apply (simp add: gpv.rel_map map_prod_def split_def) + by (rule gpv.rel_mono_strong0[of "(=)" "(=)"]) (auto simp add: gpv.rel_eq) + + + +\ \Commonly used wirings\ + +context + fixes + cnv1 :: "('icnv_usr1, 'ocnv_usr1, 'iusr1_res1 + 'iusr1_res2, 'ousr1_res1 + 'ousr1_res2) converter" and + cnv2 :: "('icnv_usr2, 'ocnv_usr2, 'iusr2_res1 + 'iusr2_res2, 'ousr2_res1 + 'ousr2_res2) converter" +begin + +\ \c1r22: a converter that has 1 interface and sends queries to two resources, + where the first and second resources have 2 and 2 interfaces respectively\ + +definition + wiring_c1r22_c1r22 :: "('icnv_usr1 + 'icnv_usr2, 'ocnv_usr1 + 'ocnv_usr2, + ('iusr1_res1 + 'iusr2_res1) + 'iusr1_res2 + 'iusr2_res2, + ('ousr1_res1 + 'ousr2_res1) + 'ousr1_res2 + 'ousr2_res2) converter" + where + "wiring_c1r22_c1r22 \ (cnv1 |\<^sub>= cnv2) \ parallel_wiring" + +end + + +\ \Special wiring converters used for the parallel composition of Fused resources\ + +definition + fused_wiring :: " + ((('iadv_core1 + 'iadv_core2) + ('iadv_rest1 + 'iadv_rest2)) + + (('iusr_core1 + 'iusr_core2) + ('iusr_rest1 + 'iusr_rest2)), + (('oadv_core1 + 'oadv_core2) + ('oadv_rest1 + 'oadv_rest2)) + + (('ousr_core1 + 'ousr_core2) + ('ousr_rest1 + 'ousr_rest2)), + (('iadv_core1 + 'iadv_rest1) + ('iusr_core1 + 'iusr_rest1)) + + (('iadv_core2 + 'iadv_rest2) + ('iusr_core2 + 'iusr_rest2)), + (('oadv_core1 + 'oadv_rest1) + ('ousr_core1 + 'ousr_rest1)) + + (('oadv_core2 + 'oadv_rest2) + ('ousr_core2 + 'ousr_rest2))) converter" + where + "fused_wiring \ (parallel_wiring |\<^sub>= parallel_wiring) \ parallel_wiring" + +definition + fused_wiring\<^sub>w + where + "fused_wiring\<^sub>w \ (parallel_wiring\<^sub>w |\<^sub>w parallel_wiring\<^sub>w) \\<^sub>w parallel_wiring\<^sub>w" + +schematic_goal + wiring_fused_wiring[wiring_intro]: "wiring ?\1 ?\2 fused_wiring fused_wiring\<^sub>w" + unfolding fused_wiring_def fused_wiring\<^sub>w_def + by(rule wiring_intro)+ + +schematic_goal WT_fused_wiring [WT_intro]: "?\1, ?\2 \\<^sub>C fused_wiring \" + unfolding fused_wiring_def + by(rule WT_intro)+ + + + +\ \Commonlu used attachments\ + +context + fixes + cnv1 :: "('icnv_usr1, 'ocnv_usr1, 'iusr1_core1 + 'iusr1_core2, 'ousr1_core1 + 'ousr1_core2) converter" and + cnv2 :: "('icnv_usr2, 'ocnv_usr2, 'iusr2_core1 + 'iusr2_core2, 'ousr2_core1 + 'ousr2_core2) converter" and + res1 :: "(('iadv_core1 + 'iadv_rest1) + ('iusr1_core1 + 'iusr2_core1) + 'iusr_rest1, + ('oadv_core1 + 'oadv_rest1) + ('ousr1_core1 + 'ousr2_core1) + 'ousr_rest1) resource" and + res2 :: "(('iadv_core2 + 'iadv_rest2) + ('iusr1_core2 + 'iusr2_core2) + 'iusr_rest2, + ('oadv_core2 + 'oadv_rest2) + ('ousr1_core2 + 'ousr2_core2) + 'ousr_rest2) resource" +begin + +\ \Attachement of two c1f22 ('f' instead of 'r' to indicate Fused Resources) converters + to two 2-interface Fused Resources, the results will be a new 2-interface Fused Resource\ + +definition + attach_c1f22_c1f22 :: "((('iadv_core1 + 'iadv_core2) + 'iadv_rest1 + 'iadv_rest2) + ('icnv_usr1 + 'icnv_usr2) + 'iusr_rest1 + 'iusr_rest2, + (('oadv_core1 + 'oadv_core2) + 'oadv_rest1 + 'oadv_rest2) + ('ocnv_usr1 + 'ocnv_usr2) + 'ousr_rest1 + 'ousr_rest2) resource" + where + "attach_c1f22_c1f22 = (((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= ((wiring_c1r22_c1r22 cnv1 cnv2) |\<^sub>= 1\<^sub>C)) \ fused_wiring) \ (res1 \ res2)" +end + + +\ \Properties of Converters attaching to Fused resources\ +context + fixes + core1 :: "('s_core1, 'e1, 'iadv_core1, 'iusr_core1, 'oadv_core1, 'ousr_core1) core" and + core2 :: "('s_core2, 'e2, 'iadv_core2, 'iusr_core2, 'oadv_core2, 'ousr_core2) core" and + rest1 :: "('s_rest1, 'e1, 'iadv_rest1, 'iusr_rest1, 'oadv_rest1, 'ousr_rest1, 'm1) rest_scheme" and + rest2 :: "('s_rest2, 'e2, 'iadv_rest2, 'iusr_rest2, 'oadv_rest2, 'ousr_rest2, 'm2) rest_scheme" +begin + +lemma parallel_oracle_fuse: + "apply_wiring fused_wiring\<^sub>w (parallel_oracle (fused_resource.fuse core1 rest1) (fused_resource.fuse core2 rest2)) = + apply_state_iso parallel_state_iso (fused_resource.fuse (parallel_core core1 core2) (parallel_rest rest1 rest2))" + supply fused_resource.fuse.simps[simp] + apply(rule ext)+ + apply(clarsimp simp add: fused_wiring\<^sub>w_def apply_state_iso_def parallel_state_iso_def parallel_wiring\<^sub>w_def) + apply(simp add: apply_wiring_def comp_wiring_def parallel2_wiring_def lassocr\<^sub>w_def swap_lassocr\<^sub>w_def rassocl\<^sub>w_def swap\<^sub>w_def) + subgoal for s_core1 s_rest1 s_core2 s_rest2 i + apply(cases "(parallel_rest rest1 rest2, ((s_core1, s_core2), (s_rest1, s_rest2)), i)" rule: fused_resource.fuse.cases) + apply(auto split!: sum.splits) + subgoal for iadv_core + by (cases iadv_core) (auto simp add: map_spmf_bind_spmf bind_map_spmf intro!: bind_spmf_cong split!: sum.splits) + subgoal for iadv_rest + by (cases iadv_rest) (auto simp add: parallel_handler_left parallel_handler_right foldl_spmf_pair_left + parallel_eoracle_def foldl_spmf_pair_right split_beta o_def map_spmf_bind_spmf bind_map_spmf) + subgoal for iusr_core + by (cases iusr_core) (auto simp add: map_spmf_bind_spmf bind_map_spmf intro!: bind_spmf_cong split!: sum.splits) + subgoal for iusr_rest + by (cases iusr_rest) (auto simp add: parallel_handler_left parallel_handler_right foldl_spmf_pair_left + parallel_eoracle_def foldl_spmf_pair_right split_beta o_def map_spmf_bind_spmf bind_map_spmf) + done + done +end + +lemma attach_callee_fuse: + "attach_callee ((cnv_adv_core \\<^sub>I cnv_adv_rest) \\<^sub>I cnv_usr_core \\<^sub>I cnv_usr_rest) (fused_resource.fuse core rest) = + apply_state_iso iso_trisplit (fused_resource.fuse (attach_core cnv_adv_core cnv_usr_core core) (attach_rest cnv_adv_rest cnv_usr_rest f_init rest))" + (is "?lhs = ?rhs") +proof(intro ext; clarify) + note fused_resource.fuse.simps [simp] + let ?tri = "\(((s11, s12), s13), (s21, s22), s23). (((s11, s21), s12, s22), s13, s23)" + fix q :: "('g + 'h) + 'i + 'j" + consider (ACore) qac where "q = Inl (Inl qac)" + | (ARest) qar where "q = Inl (Inr qar)" + | (UCore) quc where "q = Inr (Inl quc)" + | (URest) qur where "q = Inr (Inr qur)" + using fuse_callee.cases by blast + then show "?lhs (((sac, sar), (suc, sur)), (sc, sr)) q = ?rhs (((sac, sar), (suc, sur)), (sc, sr)) q" + for sac sar suc sur sc sr + proof cases + case ACore + have "map_spmf rprodl (exec_gpv (fused_resource.fuse core rest) + (left_gpv (map_gpv (map_prod Inl (\s1'. (s1', suc, sur))) id (left_gpv (map_gpv (map_prod Inl (\s1'. (s1', sar))) id (cnv_adv_core sac qac))))) + (sc, sr)) = + (map_spmf (map_prod (Inl \ Inl) (?tri \ prod.swap \ Pair ((sar, sur), sr))) + (map_spmf (\((oadv, s_adv'), s_core'). (oadv, (s_adv', suc), s_core')) + (exec_gpv (cfunc_adv core) (cnv_adv_core sac qac) sc)))" + proof(induction arbitrary: sc cnv_adv_core rule: exec_gpv_fixp_parallel_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step execl execr) + show ?case + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(subst step.IH[unfolded id_def]) + apply(simp add: spmf.map_comp o_def) + done + qed + then show ?thesis using ACore + by(simp add: apply_state_iso_def iso_trisplit_def map_spmf_conv_bind_spmf[symmetric] spmf.map_comp o_def split_def) + next + case ARest + have "bind_spmf (foldl_spmf (cpoke core) (return_spmf sc) es) (\sc'. + map_spmf rprodl (exec_gpv (fused_resource.fuse core rest) + (left_gpv (map_gpv (map_prod Inl (\s1'. (s1', suc, sur))) id (right_gpv (map_gpv (map_prod Inr (Pair sac)) id (cnv_adv_rest sar qar))))) + (sc', sr))) = + bind_spmf + (exec_gpv (\(s, es) q. map_spmf (\((out, e), s'). (out, s', es @ e)) (rfunc_adv rest s q)) (cnv_adv_rest sar qar) (sr, es)) + (map_spmf (map_prod id ?tri) \ + ((\((o_rfunc, e_lst), s_rfunc). map_spmf (\s_notify. (Inl (Inr o_rfunc), s_notify, s_rfunc)) + (map_spmf (Pair (sac, suc)) (foldl_spmf (cpoke core) (return_spmf sc) e_lst))) \ + (\((oadv, s_adv'), s_rest', es). ((oadv, es), (s_adv', sur), s_rest'))))" + for es + proof(induction arbitrary: sc sr es cnv_adv_rest rule: exec_gpv_fixp_parallel_induct) + case adm then show ?case by simp + case bottom then show ?case by simp + case (step execl execr) + show ?case + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf) + apply(subst bind_commute_spmf) + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf spmf.map_comp o_def map_spmf_conv_bind_spmf[symmetric] intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(subst bind_commute_spmf) + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf spmf.map_comp o_def map_spmf_conv_bind_spmf[symmetric] intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(simp add: bind_spmf_assoc[symmetric] bind_foldl_spmf_return foldl_spmf_append[symmetric] del: bind_spmf_assoc ) + apply(subst step.IH[unfolded id_def]) + apply(simp add: split_def o_def spmf.map_comp) + done + qed + from this[of "[]"] + show ?thesis using ARest + by(simp add: apply_state_iso_def iso_trisplit_def map_bind_spmf bind_map_spmf map_spmf_conv_bind_spmf[symmetric] foldl_spmf_pair_right) + next + case UCore + have "map_spmf rprodl (exec_gpv (fused_resource.fuse core rest) + (right_gpv (map_gpv (map_prod Inr (Pair (sac, sar))) id (left_gpv (map_gpv (map_prod Inl (\s1'. (s1', sur))) id (cnv_usr_core suc quc))))) + (sc, sr)) = + (map_spmf (map_prod (Inr \ Inl) (?tri \ prod.swap \ Pair ((sar, sur), sr))) + (map_spmf (\((ousr, s_usr'), s_core'). (ousr, (sac, s_usr'), s_core')) + (exec_gpv (cfunc_usr core) (cnv_usr_core suc quc) sc)))" + proof(induction arbitrary: sc cnv_usr_core rule: exec_gpv_fixp_parallel_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step execl execr) + show ?case + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(subst step.IH[unfolded id_def]) + apply(simp add: spmf.map_comp o_def) + done + qed + then show ?thesis using UCore + by(simp add: apply_state_iso_def iso_trisplit_def map_spmf_conv_bind_spmf[symmetric] spmf.map_comp o_def split_def) + next + case URest + have "bind_spmf (foldl_spmf (cpoke core) (return_spmf sc) es) (\sc'. + map_spmf rprodl (exec_gpv (fused_resource.fuse core rest) + (right_gpv (map_gpv (map_prod Inr (Pair (sac, sar))) id (right_gpv (map_gpv (map_prod Inr (Pair suc)) id (cnv_usr_rest sur qur))))) + (sc', sr))) = + bind_spmf + (exec_gpv (\(s, es) q. map_spmf (\((out, e), s'). (out, s', es @ e)) (rfunc_usr rest s q)) (cnv_usr_rest sur qur) (sr, es)) + (map_spmf (map_prod id ?tri) \ + ((\((o_rfunc, e_lst), s_rfunc). map_spmf (\s_notify. (Inr (Inr o_rfunc), s_notify, s_rfunc)) + (map_spmf (Pair (sac, suc)) (foldl_spmf (cpoke core) (return_spmf sc) e_lst))) \ + (\((ousr, s_usr'), s_rest', es). ((ousr, es), (sar, s_usr'), s_rest'))))" + for es + proof(induction arbitrary: sc sr es cnv_usr_rest rule: exec_gpv_fixp_parallel_induct) + case adm then show ?case by simp + case bottom then show ?case by simp + case (step execl execr) + show ?case + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf) + apply(subst bind_commute_spmf) + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf spmf.map_comp o_def map_spmf_conv_bind_spmf[symmetric] intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(subst bind_commute_spmf) + apply(clarsimp simp add: gpv.map_sel map_bind_spmf bind_map_spmf spmf.map_comp o_def map_spmf_conv_bind_spmf[symmetric] intro!: bind_spmf_cong[OF refl] split!: generat.split) + apply(simp add: bind_spmf_assoc[symmetric] bind_foldl_spmf_return foldl_spmf_append[symmetric] del: bind_spmf_assoc ) + apply(subst step.IH[unfolded id_def]) + apply(simp add: split_def o_def spmf.map_comp) + done + qed + from this[of "[]"] show ?thesis using URest + by(simp add: apply_state_iso_def iso_trisplit_def map_bind_spmf bind_map_spmf map_spmf_conv_bind_spmf[symmetric] foldl_spmf_pair_right) + qed +qed + +lemma attach_parallel_fuse': + "(CNV cnv_adv_core s_a_c |\<^sub>= CNV cnv_adv_rest s_a_r) |\<^sub>= (CNV cnv_usr_core s_u_c |\<^sub>= CNV cnv_usr_rest s_u_r) \ + RES (fused_resource.fuse core rest) (s_r_c, s_r_r) = + RES (fused_resource.fuse (attach_core cnv_adv_core cnv_usr_core core) (attach_rest cnv_adv_rest cnv_usr_rest f_init rest)) (((s_a_c, s_u_c), s_r_c), ((s_a_r, s_u_r), s_r_r))" + apply(fold conv_callee_parallel) + apply(unfold attach_CNV_RES) + apply(subst attach_callee_fuse) + apply(subst resource_of_oracle_state_iso) + apply simp + apply(simp add: iso_trisplit_def) + done + + +\ \Moving event translators from rest to the core\ + +context + fixes + einit :: 's_event and + etran :: "('s_event, 'ievent, 'oevent list) oracle'" and + rest :: "('s_rest, 'ievent, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) rest_wstate" and + core :: "('s_core, 'oevent, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" +begin + +primcorec + translate_rest :: "('s_event \ 's_rest, 'oevent, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) rest_wstate" + where + "rinit translate_rest = (einit, rinit rest)" + | "rfunc_adv translate_rest = translate_eoracle etran (extend_state_oracle (rfunc_adv rest))" + | "rfunc_usr translate_rest = translate_eoracle etran (extend_state_oracle (rfunc_usr rest))" + +primcorec + translate_core :: "('s_event \ 's_core, 'ievent, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + where + "cpoke translate_core = (\(s_event, s_core) event. + bind_spmf (etran s_event event) (\(events, s_event'). + map_spmf (\s_core'. (s_event', s_core')) (foldl_spmf (cpoke core) (return_spmf s_core) events)))" + | "cfunc_adv translate_core = extend_state_oracle (cfunc_adv core)" + | "cfunc_usr translate_core = extend_state_oracle (cfunc_usr core)" + + +lemma WT_translate_rest [WT_intro]: + assumes "WT_rest \_adv \_usr I_rest rest" + shows "WT_rest \_adv \_usr (pred_prod (\_. True) I_rest) translate_rest" + by(rule WT_rest.intros)(auto simp add: translate_eoracle_def simp add: WT_restD_rinit[OF assms] dest!: WT_restD(1,2)[OF assms]) + + +lemma fused_resource_move_translate: + "fused_resource.fuse core translate_rest = apply_state_iso iso_swapar (fused_resource.fuse translate_core rest)" +proof - + note [simp] = exec_gpv_bind spmf.map_comp o_def map_bind_spmf bind_map_spmf bind_spmf_const + + show ?thesis + apply (rule ext) + apply (rule ext) + subgoal for s query + apply (cases s) + subgoal for s_core s_event s_rest + apply (cases query) + subgoal for q_adv + apply (cases q_adv) + subgoal for q_acore + by (simp add: apply_state_iso_def iso_swapar_def fused_resource.fuse.simps split_def map_prod_def) + subgoal for q_arest + apply (simp add: apply_state_iso_def iso_swapar_def fused_resource.fuse.simps) + apply (simp add: translate_eoracle_def split_def) + apply(rule bind_spmf_cong[OF refl]) + apply(subst foldl_spmf_chain[simplified split_def]) + by simp + done + subgoal for q_usr + apply (cases q_usr) + subgoal for q_ucore + by (simp add: apply_state_iso_def iso_swapar_def fused_resource.fuse.simps split_def map_prod_def) + subgoal for q_urest + apply (simp add: apply_state_iso_def iso_swapar_def fused_resource.fuse.simps) + apply (simp add: translate_eoracle_def split_def) + apply(rule bind_spmf_cong[OF refl]) + apply(subst foldl_spmf_chain[simplified split_def]) + by simp + done + done + done + done +qed + + + +end + + +\ \Moving interfaces between rest and core\ + +lemma + fuse_ishift_core_to_rest: + assumes "cpoke core' = (\s. case_sum (\q. fn s q) (cpoke core s))" + and "cfunc_adv core = cfunc_adv core'" + and "cfunc_usr core = cfunc_usr core' \\<^sub>O (\s i. map_spmf (Pair (h_out i)) (fn s i))" + and "rfunc_adv rest' = (\s q. map_spmf (apfst (apsnd (map Inr))) (rfunc_adv rest s q))" + and "rfunc_usr rest' = plus_eoracle (\s i. return_spmf ((h_out i, [i]), s)) (rfunc_usr rest)" + shows "fused_resource.fuse core rest = apply_wiring (1\<^sub>w |\<^sub>w lassocr\<^sub>w) (fused_resource.fuse core' rest')" (is "?L = ?R") +proof - + note [simp] = fused_resource.fuse.simps apply_wiring_def lassocr\<^sub>w_def parallel2_wiring_def + plus_eoracle_def map_spmf_conv_bind_spmf map_prod_def map_fun_def split_def o_def + + have "?L s q = ?R s q" for s q + apply (cases q; cases s) + subgoal for q_adv + by (cases q_adv) (simp_all add: assms(1, 2, 4)) + subgoal for q_usr + apply (cases q_usr) + subgoal for q_usr_core + apply (cases q_usr_core) + subgoal for q_nrm + by (simp add: assms(3)) + by (simp add: assms(1, 3, 5)) + by (simp add: assms(1, 5)) + done + + then show ?thesis + by blast +qed + + +lemma move_simulator_interface: + defines "x_ifunc \ (\ifunc core (se, sc) q. do { + ((out, es), se') \ ifunc se q; + sc' \ foldl_spmf (cpoke core) (return_spmf sc) es; + return_spmf (out, se', sc') })" + assumes "cpoke core' = cpoke (translate_core etran core)" + and "cfunc_adv core' = \(cfunc_adv core) \\<^sub>O x_ifunc ifunc core" + and "cfunc_usr core' = cfunc_usr (translate_core etran core)" + and "rinit rest = (einit, rinit rest')" + and "rfunc_adv rest = (\s q. case q of + Inl ql \ map_spmf (apfst (map_prod Inl id)) ((ifunc\) s ql) + | Inr qr \ map_spmf (apfst (map_prod Inr id)) ((translate_eoracle etran (\(rfunc_adv rest'))) s qr))" + and "rfunc_usr rest = translate_eoracle etran (\(rfunc_usr rest'))" + shows "fused_resource.fuse core rest = apply_wiring (rassocl\<^sub>w |\<^sub>w (id, id)) + (apply_state_iso (rprodl o (apfst prod.swap), (apfst prod.swap) o lprodr) + (fused_resource.fuse core' rest'))" + (is "?L = ?R") +proof - + note [simp] = fused_resource.fuse.simps apply_wiring_def rassocl\<^sub>w_def parallel2_wiring_def apply_state_iso_def + exec_gpv_bind spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def split_def + + have "?L (sc, st, sr) q = ?R (sc, st, sr) q" for sc st sr q + apply (simp add: map_fun_def map_prod_def prod.swap_def apfst_def lprodr_def rprodl_def id_def) + using assms apply (cases q) + subgoal for q_adv + apply (cases q_adv) + subgoal for q_adv_core + by (simp add: map_prod_def) + subgoal for q_adv_rest + apply (cases q_adv_rest) + subgoal for q_adv_rest_ifunc + by simp + subgoal for q_adv_rest_etran + apply (simp add: translate_eoracle_def) + apply(rule bind_spmf_cong[OF refl]) + apply(subst foldl_spmf_chain[simplified split_def]) + by simp + done + done + subgoal for q_usr + apply (cases q_usr) + subgoal for q_usr_core + by (simp add: map_prod_def) + subgoal for q_usr_rest + apply (simp add: translate_eoracle_def extend_state_oracle_def) + apply(rule bind_spmf_cong[OF refl]) + apply(subst foldl_spmf_chain[simplified split_def]) + by simp + done + done + + then show ?thesis + by force +qed + + +end diff --git a/thys/Constructive_Cryptography_CM/Constructions/DH_OTP.thy b/thys/Constructive_Cryptography_CM/Constructions/DH_OTP.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Constructions/DH_OTP.thy @@ -0,0 +1,102 @@ +theory DH_OTP imports + One_Time_Pad + Diffie_Hellman_CC +begin + +text \ + We need both a group structure and a boolean algebra. + Unfortunately, records allow only one extension slot, so we can't have just a single + structure with both operations. +\ + +context diffie_hellman begin + +lemma WT_ideal_rest [WT_intro]: + assumes WT_auth1_rest [WT_intro]: "WT_rest \_adv_rest1 \_usr_rest1 I_auth1_rest auth1_rest" + and WT_auth2_rest [WT_intro]: "WT_rest \_adv_rest2 \_usr_rest2 I_auth2_rest auth2_rest" + shows "WT_rest (\_full \\<^sub>\ (\_adv_rest1 \\<^sub>\ \_adv_rest2)) ((\_full \\<^sub>\ \_full) \\<^sub>\ (\_usr_rest1 \\<^sub>\ \_usr_rest2)) + (\(_, s). pred_prod I_auth1_rest I_auth2_rest s) (ideal_rest auth1_rest auth2_rest)" + apply(rule WT_rest.intros) + subgoal + by(auto 4 4 split: sum.splits simp add: translate_eoracle_def parallel_eoracle_def dest: assms[THEN WT_restD_rfunc_adv]) + subgoal + apply(auto 4 4 split: sum.splits simp add: translate_eoracle_def parallel_eoracle_def plus_eoracle_def dest: assms[THEN WT_restD_rfunc_usr]) + apply(simp add: map_sum_def split: sum.splits) + done + subgoal by(simp add: assms[THEN WT_restD_rinit]) +done + +end + + +locale dh_otp = dh: diffie_hellman \ + otp: one_time_pad \ + for \ :: "'grp cyclic_group" + and \ :: "'grp boolean_algebra" + + assumes carrier_\_\: "carrier \ = carrier \" +begin + +theorem secure: + assumes "WT_rest \_adv_resta \_usr_resta I_auth_rest auth_rest" + and "WT_rest \_adv_rest1 \_usr_rest1 I_auth1_rest auth1_rest" + and "WT_rest \_adv_rest2 \_usr_rest2 I_auth2_rest auth2_rest" + shows + "constructive_security_obsf + (1\<^sub>C |\<^sub>= wiring_c1r22_c1r22 (CNV otp.enc_callee ()) (CNV otp.dec_callee ()) |\<^sub>= 1\<^sub>C \ + fused_wiring \ diffie_hellman.real_resource \ auth1_rest auth2_rest \ dh.auth.resource auth_rest) + (otp.sec.resource (otp.ideal_rest (dh.ideal_rest auth1_rest auth2_rest) auth_rest)) + ((1\<^sub>C \ + (parallel_wiring \ ((let sim = CNV dh.sim_callee None in (sim |\<^sub>= 1\<^sub>C) \ lassocr\<^sub>C) |\<^sub>= 1\<^sub>C) \ parallel_wiring) \ + 1\<^sub>C) \ + (otp.sim |\<^sub>= 1\<^sub>C)) + ((((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (otp.sec.Inp_Fedit ` carrier \) UNIV)) \\<^sub>\ + (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (otp.sec.Inp_Fedit ` carrier \) UNIV))) \\<^sub>\ + (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (otp.sec.Inp_Fedit ` carrier \) UNIV))) \\<^sub>\ + ((\_adv_rest1 \\<^sub>\ \_adv_rest2) \\<^sub>\ \_adv_resta)) + ((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (otp.sec.Inp_Fedit ` carrier \) UNIV)) \\<^sub>\ + ((\_full \\<^sub>\ (\_adv_rest1 \\<^sub>\ \_adv_rest2)) \\<^sub>\ \_adv_resta)) + ((\_uniform (otp.sec.Inp_Send ` carrier \) UNIV \\<^sub>\ \_uniform UNIV (otp.sec.Out_Recv ` carrier \)) \\<^sub>\ + (((\_full \\<^sub>\ \_full) \\<^sub>\ (\_usr_rest1 \\<^sub>\ \_usr_rest2)) \\<^sub>\ \_usr_resta)) + \ (0 + (ddh.advantage \ + (diffie_hellman.DH_adversary \ auth1_rest auth2_rest + (absorb + (absorb \ + (obsf_converter (1\<^sub>C |\<^sub>= wiring_c1r22_c1r22 (CNV otp.enc_callee ()) (CNV otp.dec_callee ()) |\<^sub>= 1\<^sub>C))) + (obsf_converter + (fused_wiring \ (1\<^sub>C |\<^sub>\ converter_of_resource (1\<^sub>C |\<^sub>= 1\<^sub>C \ dh.auth.resource auth_rest)))))) + + 0))" + using assms apply - + apply(rule constructive_security_obsf_composability) + apply(rule otp.secure) + apply(rule WT_intro, assumption+) + unfolding otp.real_resource_def attach_c1f22_c1f22_def[abs_def] attach_compose + apply(rule constructive_security_obsf_lifting_[where w_adv_real="1\<^sub>C" and w_adv_ideal_inv="1\<^sub>C"]) + apply(rule parallel_constructive_security_obsf_fuse) + apply(fold carrier_\_\)[1] + apply(rule dh.secure, assumption, assumption, rule constructive_security_obsf_trivial) + defer + defer + defer + apply(rule WT_intro)+ + apply(simp add: comp_converter_id_left) + apply(rule parallel_converter2_id_id pfinite_intro wiring_intro)+ + apply(rule WT_intro|assumption)+ + apply simp + apply(unfold wiring_c1r22_c1r22_def) + apply(rule WT_intro)+ + apply(fold carrier_\_\)[1] + apply(rule WT_intro)+ + + apply(rule pfinite_intro) + apply(rule pfinite_intro) + apply(rule pfinite_intro) + apply(rule pfinite_intro) + apply(rule pfinite_intro) + apply(unfold carrier_\_\) + apply(rule pfinite_intro) + apply(rule WT_intro)+ + apply(rule pfinite_intro) + done + +end + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Constructions/Diffie_Hellman_CC.thy b/thys/Constructive_Cryptography_CM/Constructions/Diffie_Hellman_CC.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Constructions/Diffie_Hellman_CC.thy @@ -0,0 +1,2368 @@ +theory Diffie_Hellman_CC + imports + Game_Based_Crypto.Diffie_Hellman + "../Asymptotic_Security" + "../Construction_Utility" + "../Specifications/Key" + "../Specifications/Channel" +begin + +hide_const (open) Resumption.Pause Monomorphic_Monad.Pause Monomorphic_Monad.Done + +no_notation Sublist.parallel (infixl "\" 50) +no_notation plus_oracle (infix "\\<^sub>O" 500) + + +section \Diffie-Hellman construction\ + +locale diffie_hellman = + auth: ideal_channel "id :: 'grp \ 'grp" False + + key: ideal_key "carrier \" + + cyclic_group \ + for + \ :: "'grp cyclic_group" (structure) +begin + + +subsection \Defining user callees\ + +datatype 'grp' cstate = CState_Void | CState_Half nat | CState_Full "nat \ 'grp'" + +datatype icnv_alice = Inp_Activation_Alice +datatype icnv_bob = Iact_Activation_Bob + +datatype ocnv_alice = Out_Activation_Alice +datatype ocnv_bob = Out_Activation_Bob + +fun alice_callee :: "'grp cstate \ key.iusr_alice + icnv_alice + \ (('grp key.ousr_alice + ocnv_alice) \ 'grp cstate, 'grp auth.iusr_alice + auth.iusr_bob, auth.ousr_alice + 'grp auth.ousr_bob) gpv" + where + "alice_callee CState_Void (Inr _) = do { + x \ lift_spmf (sample_uniform (order \)); + let msg = \<^bold>g [^] x; + Pause + (Inl (auth.Inp_Send msg)) + (\rsp. case rsp of + Inl _ \ Done (Inr Out_Activation_Alice, CState_Half x) + | Inr _ \ Fail) }" + | "alice_callee (CState_Half x) (Inl _) = + Pause + (Inr auth.Inp_Recv) + (\rsp. case rsp of + Inl _ \ Fail + | Inr msg \ case msg of + auth.Out_Recv gy \ + let key = gy [^] x in + Done (Inl (key.Out_Alice key), CState_Full (x, key)))" + | "alice_callee (CState_Full (x, key)) (Inl _) = Done (Inl (key.Out_Alice key), CState_Full (x, key))" + | "alice_callee _ _ = Fail" + +fun bob_callee :: "'grp cstate \ key.iusr_bob + icnv_bob + \ (('grp key.ousr_bob + ocnv_bob) \ 'grp cstate, auth.iusr_bob + 'grp auth.iusr_alice, 'grp auth.ousr_bob + auth.ousr_alice) gpv" + where + "bob_callee CState_Void (Inr _) = do { + y \ lift_spmf (sample_uniform (order \)); + let msg = \<^bold>g [^] y; + Pause + (Inr (auth.Inp_Send msg)) + (\rsp. case rsp of + Inl _ \ Fail + | Inr _ \ Done (Inr Out_Activation_Bob, CState_Half y)) }" + | "bob_callee (CState_Half y) (Inl _) = + Pause + (Inl auth.Inp_Recv) + (\rsp. case rsp of + Inl msg \ case msg of + auth.Out_Recv gx \ + let k = gx [^] y in + Done (Inl (key.Out_Bob k), CState_Full (y, k)) + | Inr _ \ Fail)" + | "bob_callee (CState_Full (y, key)) (Inl _) = Done (Inl (key.Out_Bob key), CState_Full (y, key))" + | "bob_callee _ _ = Fail" + + +subsection \Defining adversary callee\ +(* +astate_rbase \ ranmed to \ asate_chann +datatype astate_chann = AState_Available | AState_Unavailable +type_synonym 'grp' astate_lbase = "('grp' \ 'grp') option" +type_synonym 'grp' astate_single = "'grp' astate_lbase \ astate_rbase"*) +type_synonym 'grp' astate = "('grp' \ 'grp') option" + +type_synonym 'grp' isim = "'grp' auth.iadv + 'grp' auth.iadv" +datatype osim = Out_Simulator + +fun sim_callee_base :: "(('grp \ 'grp) \ 'grp ) \ ('grp astate, 'grp auth.iadv, 'grp auth.oadv) oracle'" + where + "sim_callee_base _ _ (Inl _) = return_pmf None" + | "sim_callee_base pick gpair_opt (Inr (Inl _)) = do { + sample \ do { + x \ sample_uniform (order \); + y \ sample_uniform (order \); + return_spmf (\<^bold>g [^] x, \<^bold>g [^] y) }; + let sample' = case_option sample id gpair_opt; + return_spmf (Inr (Inl (auth.Out_Look (pick sample'))), Some sample') }" + | "sim_callee_base _ gpair_opt (Inr (Inr _ )) = return_spmf (Inr (Inr auth.Out_Fedit), gpair_opt)" + +fun sim_callee :: "'grp astate \ 'grp auth.iadv + 'grp auth.iadv + \ (('grp auth.oadv + 'grp auth.oadv) \ 'grp astate, key.iadv + 'grp isim, key.oadv + osim) gpv" + where + "sim_callee s_gpair query = + (let handle = (\gpair_pick wrap_out q_split. do { + _ \ Pause (Inr query) Done; + (out, s_gpair') \ lift_spmf (sim_callee_base gpair_pick s_gpair q_split); + Done (wrap_out out, s_gpair') }) in + case_sum (handle fst Inl) (handle snd Inr) query)" + +subsection \Defining event-translator\ + +datatype estate_base = EState_Void | EState_Store | EState_Collect +type_synonym estate = "bool \ (estate_base \ auth.s_shell) \ estate_base \ auth.s_shell" + +definition einit :: estate + where + "einit \ (False, (EState_Void, {}), EState_Void, {})" + +definition eleak :: "(estate, key.event, 'grp isim, osim) eoracle" + where + "eleak \ (\(s_flg, (s_event1, s_shell1), s_event2, s_shell2) query. + let handle_arg1 = (\s q. case (s, q) of (EState_Store, Some (Inr (Inr _))) \ (True, EState_Collect) | (s', _) \ (False, s')) in + let handle_arg2 = (\s q D. case (s, q) of (EState_Store, Inr _) \ D | _ \ return_pmf None) in + let (is_ch1, s_event1') = handle_arg1 s_event1 (case_sum Some (\_. None) query) in + let (is_ch2, s_event2') = handle_arg1 s_event2 (case_sum (\_. None) Some query) in + let check_pst1 = is_ch1 \ s_event2' \ EState_Void \ auth.Bob \ s_shell1 \ auth.Alice \ s_shell2 in + let check_pst2 = is_ch2 \ s_event1' \ EState_Void \ auth.Alice \ s_shell1 \ auth.Bob \ s_shell2 in + let e_pstfix1 = if check_pst1 then [key.Event_Shell key.Bob] else [] in + let e_pstfix2 = if check_pst2 then [key.Event_Shell key.Alice] else [] in + let e_prefix = if \s_flg then [key.Event_Kernel] else [] in + let (s_flg', event) = if is_ch2 \ is_ch1 then (True, e_prefix @ e_pstfix1 @ e_pstfix2) else (s_flg, []) in + let result_base = return_spmf ((Out_Simulator, event), s_flg', (s_event1', s_shell1), s_event2', s_shell2) in + case_sum (handle_arg2 s_event1) (handle_arg2 s_event2) query result_base)" + +fun etran_base :: "(key.party \ key.party \ key.party \ key.party) + \ (estate, auth.event, key.event list) oracle'" + where + "etran_base mod_event (s_flg, (s_event1, s_shell1), s_event2, s_shell2) (auth.Event_Shell party) = + (let party_dual = auth.case_party (auth.Bob) (auth.Alice) party in + let epair = auth.case_party prod.swap id party (key.Bob, key.Alice) in + let (s_event_eq, s_event_neq) = auth.case_party prod.swap id party (s_event1, s_event2) in + let check = party_dual \ s_shell2 \ s_event_eq = EState_Collect \ s_event_neq \ EState_Void in + let event = if check then [key.Event_Shell ((fst o mod_event) epair)] else [] in + let s_shell1' = insert party s_shell1 in + if party \ s_shell1 then + return_pmf None + else + return_spmf (event, s_flg, (s_event1, s_shell1'), s_event2, s_shell2))" + +fun etran :: "(estate, (icnv_alice + icnv_bob) + auth.event + auth.event, key.event list) oracle'" + where + "etran (s_flg, (EState_Void, s_shell1), s_event2, s_shell2) (Inl (Inl _)) = + (let check = (s_event2 = EState_Collect \ auth.Alice \ s_shell1 \ auth.Bob \ s_shell2) in + let event = if check then [key.Event_Shell key.Alice] else [] in + let state = (s_flg, (EState_Store, s_shell1), s_event2, s_shell2) in + if auth.Alice \ s_shell1 then return_spmf (event, state) else return_pmf None)" + | "etran (s_flg, (s_event1, s_shell1), EState_Void, s_shell2) (Inl (Inr _)) = + (let check = (s_event1 = EState_Collect \ auth.Bob \ s_shell1 \ auth.Alice \ s_shell2) in + let event = if check then [key.Event_Shell key.Bob] else [] in + let state = (s_flg, (s_event1, s_shell1), EState_Store, s_shell2) in + if auth.Alice \ s_shell2 then return_spmf (event, state) else return_pmf None)" + | "etran state (Inr query) = + (let handle = (\mod_s mod_e q. do { + (evt, state') \ etran_base mod_e (mod_s state) q; + return_spmf (evt, mod_s state') }) in + case_sum (handle id id) (handle (apsnd prod.swap) prod.swap) query)" + | "etran _ _ = return_pmf None" + + +subsection \Defining Ideal and Real constructions\ + +context + fixes + auth1_rest :: "('auth1_s_rest, auth.event, 'auth1_iadv_rest, 'auth1_iusr_rest, 'auth1_oadv_rest, 'auth1_ousr_rest) rest_wstate" and + auth2_rest :: "('auth2_s_rest, auth.event, 'auth2_iadv_rest, 'auth2_iusr_rest, 'auth2_oadv_rest, 'auth2_ousr_rest) rest_wstate" +begin + +primcorec ideal_core_alt + where + "cpoke ideal_core_alt = cpoke (translate_core etran key.core)" + | "cfunc_adv ideal_core_alt = \(cfunc_adv key.core) \\<^sub>O (\(se, sc) q. do { + ((out, es), se') \ eleak se q; + sc' \ foldl_spmf (cpoke key.core) (return_spmf sc) es; + return_spmf (out, se', sc') })" + | "cfunc_usr ideal_core_alt = cfunc_usr (translate_core etran key.core)" + +primcorec ideal_rest_alt + where + "rinit ideal_rest_alt = rinit (parallel_rest auth1_rest auth2_rest)" + | "rfunc_adv ideal_rest_alt = (\s q. map_spmf (apfst (apsnd (map Inr))) (rfunc_adv (parallel_rest auth1_rest auth2_rest) s q))" + | "rfunc_usr ideal_rest_alt = ( + let handle = map_sum (\_ :: icnv_alice. Out_Activation_Alice) (\_ :: icnv_bob. Out_Activation_Bob) in + plus_eoracle (\s q. return_spmf ((handle q, [q]), s)) (rfunc_usr (parallel_rest auth1_rest auth2_rest)))" + +primcorec ideal_rest + where + "rinit ideal_rest = (einit, rinit ideal_rest_alt)" + | "rfunc_adv ideal_rest = (\s q. case q of + Inl ql \ map_spmf (apfst (map_prod Inl id)) (eleak\ s ql) + | Inr qr \ map_spmf (apfst (map_prod Inr id)) (translate_eoracle etran \(rfunc_adv ideal_rest_alt) s qr))" + | "rfunc_usr ideal_rest = translate_eoracle etran \(rfunc_usr ideal_rest_alt)" + +definition ideal_resource + where + "ideal_resource \ + (let sim = CNV sim_callee None in + attach ((sim |\<^sub>= 1\<^sub>C ) \ lassocr\<^sub>C |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C) (key.resource ideal_rest))" + +definition real_resource + where + "real_resource \ + (let dh_wiring = parallel_wiring \ (CNV alice_callee CState_Void |\<^sub>= CNV bob_callee CState_Void) \ parallel_wiring \ (1\<^sub>C |\<^sub>= swap\<^sub>C) in + attach (((1\<^sub>C |\<^sub>= 1\<^sub>C) |\<^sub>= rassocl\<^sub>C \ (dh_wiring |\<^sub>= 1\<^sub>C)) \ fused_wiring) ((auth.resource auth1_rest) \ (auth.resource auth2_rest)))" + + +subsection \Wiring and simplifying the Ideal construction\ + +abbreviation basic_rest_sinit + where + "basic_rest_sinit \ (((), ()), rinit auth1_rest, rinit auth2_rest)" + +primcorec basic_rest :: "((unit \ unit) \ _, _, _, _, _, _, _) rest_scheme" + where + "rinit basic_rest = (rinit auth1_rest, rinit auth2_rest)" + | "rfunc_adv basic_rest = \(parallel_eoracle (rfunc_adv auth1_rest) (rfunc_adv auth2_rest))" + | "rfunc_usr basic_rest = \(parallel_eoracle (rfunc_usr auth1_rest) (rfunc_usr auth2_rest))" + +definition ideal_s_core' :: "('grp astate \ _) \ _ \ 'grp key.state" + where + "ideal_s_core' \ ((None, ()), einit, key.PState_Store, {})" + +definition ideal_s_rest' + where + "ideal_s_rest' \ basic_rest_sinit" + +primcorec ideal_core' + where + "cpoke ideal_core' = (\(s_cnv, s_event, s_core) event. do { + (events, s_event') \ (etran s_event event); + s_core' \ foldl_spmf key.poke (return_spmf s_core) events; + return_spmf (s_cnv, s_event', s_core') })" + | "cfunc_adv ideal_core' = (\((s_sim, _), s_event_core) q. + map_spmf + (\((out, s_sim'), s_event_core'). (out, (s_sim', ()), s_event_core')) + (exec_gpv + (\key.iface_adv \\<^sub>O (\(se, sc) isim. do { + ((out, es), se') \ eleak se isim; + sc' \ foldl_spmf (cpoke key.core) (return_spmf sc) es; + return_spmf (out, se', sc') })) + (sim_callee s_sim q) s_event_core))" + | "cfunc_usr ideal_core' = (\(s_cnv, s_core) q. + map_spmf (\(out, s_core'). (out, s_cnv, s_core')) (\key.iface_usr s_core q))" + +primcorec ideal_rest' + where + "rinit ideal_rest' = rinit basic_rest" + | "rfunc_adv ideal_rest' = (\s q. map_spmf (apfst (apsnd (map Inr))) (rfunc_adv basic_rest s q))" + | "rfunc_usr ideal_rest' = ( + let handle = map_sum (\_ :: icnv_alice. Out_Activation_Alice) (\_ :: icnv_bob. Out_Activation_Bob) in + plus_eoracle (\s q. return_spmf ((handle q, [q]), s)) (rfunc_usr basic_rest))" + + +subsubsection \The ideal attachment lemma\ + +context +begin + +lemma ideal_resource_shift_interface: "key.resource ideal_rest = RES + (apply_wiring (rassocl\<^sub>w |\<^sub>w (id, id)) (fused_resource.fuse ideal_core_alt ideal_rest_alt)) + ((einit, key.PState_Store, {}), rinit ideal_rest_alt)" +proof - + have "state_iso (rprodl \ apfst prod.swap, apfst prod.swap \ lprodr)" + by(simp add: state_iso_def rprodl_def lprodr_def apfst_def; unfold_locales; simp add: split_def) + + note f1= resource_of_oracle_state_iso[OF this] + + have f2: "key.fuse ideal_rest = apply_state_iso (rprodl \ apfst prod.swap, apfst prod.swap \ lprodr) + (apply_wiring (rassocl\<^sub>w |\<^sub>w (id, id)) (fused_resource.fuse ideal_core_alt ideal_rest_alt))" + by (rule move_simulator_interface[unfolded apply_wiring_state_iso_assoc, where etran=etran and ifunc=eleak and einit=einit and + core=key.core and rest=ideal_rest and core'=ideal_core_alt and rest'=ideal_rest_alt]) simp_all + + show ?thesis + unfolding key.resource_def + by (subst f2, subst f1) simp +qed + +private lemma ideal_resource_alt_def: "ideal_resource = + (let sim = CNV sim_callee None in + let s_init = ((einit, key.PState_Store, {}), rinit ideal_rest_alt) in + attach ((sim |\<^sub>= 1\<^sub>C) |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C) (RES (fused_resource.fuse ideal_core_alt ideal_rest_alt) s_init))" +proof - + note ideal_resource_shift_interface + moreover have "sim = CNV sim_callee None \ + (sim |\<^sub>= 1\<^sub>C) \ lassocr\<^sub>C |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C \ RES (apply_wiring (rassocl\<^sub>w |\<^sub>w (id, id)) (fused_resource.fuse ideal_core_alt ideal_rest_alt)) s = + (sim |\<^sub>= 1\<^sub>C) |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C \ RES (fused_resource.fuse ideal_core_alt ideal_rest_alt) s" (is "?L \ ?R") for sim s + proof - + have fact1: "\_full, \_full \\<^sub>\ \_full \\<^sub>C CNV sim_callee s \" for s + apply(subst WT_converter_of_callee) + apply (rule WT_gpv_\_mono) + apply (rule WT_gpv_full) + apply (rule \_full_le_plus_\) + apply(rule order_refl) + apply(rule order_refl) + by (simp_all add: ) + + have fact2: "(\_full \\<^sub>\ (\_full \\<^sub>\ \_full)) \\<^sub>\ (\_full \\<^sub>\ \_full) \c + apply_wiring (rassocl\<^sub>w |\<^sub>w (id, id)) (fused_resource.fuse ideal_core_alt ideal_rest_alt) s \" for s + apply (rule WT_calleeI) + subgoal for call + apply (cases s, cases call) + apply (rename_tac [!] x) + apply (case_tac [!] x) + apply (rename_tac [2] y) + apply (case_tac [2] y) + by (auto simp add: apply_wiring_def rassocl\<^sub>w_def parallel2_wiring_def fused_resource.fuse.simps) + done + + note [simp] = spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def + + assume ?L + then show ?R + apply simp + apply (subst (1 2) conv_callee_parallel_id_right[symmetric, where s'="()"]) + apply(subst eq_resource_on_UNIV_iff[symmetric]) + apply (subst eq_resource_on_trans) + apply (rule eq_\_attach_on') + defer + apply (rule parallel_converter2_eq_\_cong) + apply (rule comp_converter_of_callee_wiring) + apply (rule wiring_lassocr) + apply (subst conv_callee_parallel_id_right) + apply(rule WT_intro)+ + apply (rule fact1) + apply(rule WT_intro)+ + apply (rule eq_\_converter_reflI) + apply(rule WT_intro)+ + defer + apply (subst (1 2 3 4) converter_of_callee_id_oracle[symmetric, where s="()"]) + apply (subst conv_callee_parallel[symmetric])+ + apply (subst (1 2) attach_CNV_RES) + subgoal + apply (rule eq_resource_on_resource_of_oracleI[where S="(=)"]) + defer + apply simp + apply (rule rel_funI)+ + apply (simp add: prod.rel_eq eq_on_def) + subgoal for s' s q' q + apply (cases s; cases q) + apply (rename_tac [!] x) + apply (case_tac [!] x) + apply (rename_tac [!] y) + apply (case_tac [4] y) + apply (auto simp add: apply_wiring_def parallel2_wiring_def attach_wiring_right_def + rassocl\<^sub>w_def lassocr\<^sub>w_def map_fun_def map_prod_def split_def) + subgoal for s_flg _ _ _ _ _ _ _ _ _ q + apply (case_tac "(s_flg, q)" rule: sim_callee.cases) + apply (simp_all add: split_def split!: sum.split if_splits cong: if_cong) + by (rule rel_spmf_bindI'[where A="(=)"], simp, clarsimp split!: sum.split if_splits + simp add: split_def map_gpv_conv_bind[symmetric] map_lift_spmf map'_lift_spmf)+ + by (simp add: spmf_rel_eq map_fun_def id_oracle_def split_def; + rule bind_spmf_cong[OF refl], clarsimp split!: sum.split if_splits + simp add: split_def map_gpv_conv_bind[symmetric] map_lift_spmf map'_lift_spmf)+ + done + apply simp + apply (rule WT_resource_of_oracle[OF fact2]) + by simp + qed + + ultimately show ?thesis + unfolding ideal_resource_def by simp +qed + +lemma attach_ideal: "ideal_resource = RES (fused_resource.fuse ideal_core' ideal_rest') (ideal_s_core', ideal_s_rest')" +proof - + + have fact1: "ideal_rest' = attach_rest 1\<^sub>I 1\<^sub>I id ideal_rest_alt" (is "?L = ?R") + proof - + note [simp] = spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def + + have "rinit ?L = rinit ?R" + by simp + + moreover have "rfunc_adv ?L = rfunc_adv ?R" + unfolding attach_rest_id_oracle_adv + by (simp add: extend_state_oracle_def split_def map_spmf_conv_bind_spmf) + + moreover have "rfunc_usr ?L = rfunc_usr ?R" + unfolding attach_rest_id_oracle_usr + apply (rule ext)+ + subgoal for s q by (cases q) (simp_all add: split_def extend_state_oracle_def plus_eoracle_def) + done + + ultimately show ?thesis + by (coinduction) simp + qed + + have fact2: "ideal_core' = attach_core sim_callee 1\<^sub>I ideal_core_alt" (is "?L = ?R") + proof - + + have "cpoke ?L = cpoke ?R" + by (simp add: split_def map_spmf_conv_bind_spmf) + + moreover have "cfunc_adv ?L = cfunc_adv ?R" + unfolding attach_core_def + by (simp add: split_def) + + moreover have "cfunc_usr ?L = cfunc_usr ?R" + unfolding attach_core_id_oracle_usr + by simp + + ultimately show ?thesis + by (coinduction) simp + qed + + show ?thesis + unfolding ideal_resource_alt_def Let_def + apply(subst (1 2 3) converter_of_callee_id_oracle[symmetric, of "()"]) + apply(subst attach_parallel_fuse') + by (simp add: fact1 fact2 ideal_s_core'_def ideal_s_rest'_def) +qed + +end + + +subsection \Wiring and simplifying the Real construction\ + +definition real_s_core' :: "(_ \ 'grp cstate \ 'grp cstate) \ 'grp auth.state \ 'grp auth.state" + where + "real_s_core' \ (((), CState_Void, CState_Void), (auth.State_Void, {}), (auth.State_Void, {}))" + +definition real_s_rest' + where + "real_s_rest' \ basic_rest_sinit" + +primcorec real_core' :: "((unit \ _) \ _, _, _, _, _, _) core" + where + "cpoke real_core' = (\(s_advusr, s_core) event. + map_spmf (Pair s_advusr) (parallel_handler auth.poke auth.poke s_core event))" + | "cfunc_adv real_core' = \(auth.iface_adv \\<^sub>O auth.iface_adv)" + | "cfunc_usr real_core' = (\((s_adv, s_usr), s_core) iusr. + let handle_req = lsumr \ map_sum id (rsuml \ map_sum swap_sum id \ lsumr) \ rsuml in + let handle_ret = lsumr \ (map_sum id (rsuml \ (map_sum swap_sum id \ lsumr)) \ rsuml) \ map_sum id swap_sum in + let handle_inp = map_sum id swap_sum \ (lsumr \ map_sum id (rsuml \ map_sum swap_sum id \ lsumr) \ rsuml) in + let handle_out = apfst (lsumr \ (map_sum id (rsuml \ (map_sum swap_sum id \ lsumr)) \ rsuml)) in + map_spmf + (\((ousr, s_usr'), s_core'). (ousr, (s_adv, s_usr'), s_core')) + (exec_gpv + (auth.iface_usr \\<^sub>O auth.iface_usr) + (map_gpv' + handle_out handle_inp handle_ret + ((alice_callee \\<^sub>I bob_callee) s_usr (handle_req iusr))) + s_core))" + +definition real_rest' :: "((unit \ unit) \ _, _, _, _, _, _, _) rest_scheme" + where + "real_rest' \ basic_rest" + + +subsubsection \The real attachment lemma\ + +lemma attach_real: "real_resource = 1\<^sub>C |\<^sub>= rassocl\<^sub>C \ RES (fused_resource.fuse real_core' real_rest') (real_s_core', real_s_rest')" +proof - + + have att_core: "real_core' = attach_core 1\<^sub>I + (attach_wiring parallel_wiring\<^sub>w + (attach_wiring_right (parallel_wiring\<^sub>w \\<^sub>w (id, id) |\<^sub>w swap\<^sub>w) (alice_callee \\<^sub>I bob_callee))) + (parallel_core auth.core auth.core)" (is "?L = ?R") + proof - + + have "cpoke ?L = cpoke ?R" + by simp + + moreover have "cfunc_adv ?L = cfunc_adv ?R" + unfolding attach_core_id_oracle_adv + by (simp add: extend_state_oracle_def) + + moreover have "cfunc_usr ?L = cfunc_usr ?R" + unfolding parallel_wiring\<^sub>w_def swap_lassocr\<^sub>w_def swap\<^sub>w_def lassocr\<^sub>w_def rassocl\<^sub>w_def + apply (simp add: parallel2_wiring_simps comp_wiring_simps) + apply (simp add: attach_wiring_simps attach_wiring_right_simps) + by (simp add: map_gpv_conv_map_gpv' map_gpv'_comp apfst_def) + + ultimately show ?thesis + by (coinduction) blast + qed + + have att_rest: "real_rest' = attach_rest 1\<^sub>I 1\<^sub>I id (parallel_rest auth1_rest auth2_rest)" (is "?L = ?R") + proof - + have "rinit ?L = rinit ?R" + unfolding real_rest'_def + by simp + + moreover have "rfunc_adv ?L = rfunc_adv ?R" + unfolding real_rest'_def attach_rest_id_oracle_adv + by (simp add: extend_state_oracle_def) + + moreover have "rfunc_usr ?L = rfunc_usr ?R" + unfolding real_rest'_def attach_rest_id_oracle_usr + by (simp add: extend_state_oracle_def) + + ultimately show ?thesis + by (coinduction) blast + qed + + have fact1: " + (\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full), (\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full) \\<^sub>C + CNV (alice_callee \\<^sub>I bob_callee) (CState_Void, CState_Void) \" + apply(subst conv_callee_parallel) + apply(rule WT_intro) + apply (rule WT_converter_of_callee[where \="\_full \\<^sub>\ \_full" and \'="\_full \\<^sub>\ \_full"]) + apply (rule WT_gpv_\_mono) + apply (rule WT_gpv_full) + apply (rule \_full_le_plus_\) + apply(rule order_refl) + apply(rule order_refl) + subgoal for s q + apply (cases s; cases q) + apply (auto simp add: Let_def split!: cstate.splits if_splits auth.ousr_bob.splits) + by (metis auth.ousr_bob.exhaust range_eqI) + apply (rule WT_converter_of_callee[where \="\_full \\<^sub>\ \_full" and \'="\_full \\<^sub>\ \_full"]) + apply (rule WT_gpv_\_mono) + apply (rule WT_gpv_full) + apply (rule \_full_le_plus_\) + apply(rule order_refl) + apply(rule order_refl) + subgoal for s q + apply (cases s; cases q) + apply (auto simp add: Let_def split!: cstate.splits if_splits auth.ousr_bob.splits) + by (metis auth.ousr_bob.exhaust range_eqI) + done + + + have fact2: " + (\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full),(\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full) \\<^sub>C + CNV (alice_callee \\<^sub>I bob_callee) (CState_Void, CState_Void) \ parallel_wiring \ (1\<^sub>C |\<^sub>= swap\<^sub>C) \" + apply(rule WT_intro) + apply (rule fact1) + apply(rule WT_intro)+ + done + + have fact3: " + (\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full),(\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full) \\<^sub>C + CNV (alice_callee \\<^sub>I bob_callee) (CState_Void, CState_Void) \ parallel_wiring \ (1\<^sub>C |\<^sub>= swap\<^sub>C) \ + CNV (attach_wiring_right (parallel_wiring\<^sub>w \\<^sub>w (id, id) |\<^sub>w swap\<^sub>w) (alice_callee \\<^sub>I bob_callee)) (CState_Void, CState_Void)" + apply (rule comp_converter_of_callee_wiring) + apply(rule wiring_intro)+ + apply(rule fact1) + done + + have fact4: " + (\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full),(\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full) \\<^sub>C + parallel_wiring \ CNV (alice_callee \\<^sub>I bob_callee) (CState_Void, CState_Void) \ parallel_wiring \ (1\<^sub>C |\<^sub>= swap\<^sub>C) \ + CNV (attach_wiring parallel_wiring\<^sub>w (attach_wiring_right (parallel_wiring\<^sub>w \\<^sub>w (id, id) |\<^sub>w swap\<^sub>w) (alice_callee \\<^sub>I bob_callee))) (CState_Void, CState_Void)" + apply (rule eq_\_converter_trans) + apply (rule eq_\_comp_cong) + apply(rule eq_\_converter_reflI) + apply(rule WT_intro) + apply (rule fact3) + apply (rule comp_wiring_converter_of_callee) + apply (rule wiring_intro) + apply (subst eq_\_converterD_WT[OF fact3, simplified fact2, symmetric]) + by blast + + show ?thesis + unfolding real_resource_def auth.resource_def + apply (subst resource_of_parallel_oracle[symmetric]) + apply (subst attach_compose) + apply(subst attach_wiring_resource_of_oracle) + apply (rule wiring_intro) + apply (rule WT_resource_of_oracle[where \="((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full)) \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full))"]) + subgoal for _ s + apply (rule WT_calleeI) + apply (cases s) + apply(case_tac call) + apply(rename_tac [!] x) + apply(case_tac [!] x) + apply(rename_tac [!] y) + apply(case_tac [!] y) + apply(auto simp add: fused_resource.fuse.simps) + done + apply simp + subgoal + apply (subst parallel_oracle_fuse) + apply(subst resource_of_oracle_state_iso) + apply simp + apply(simp add: parallel_state_iso_def) + apply(subst parallel_converter2_comp2_out) + apply(subst conv_callee_parallel[symmetric]) + apply(subst eq_resource_on_UNIV_iff[symmetric]) + apply(rule eq_resource_on_trans) + apply(rule eq_\_attach_on') + prefer 2 + apply (rule eq_\_comp_cong) + apply(rule eq_\_converter_reflI) + apply(rule WT_intro)+ + apply(rule parallel_converter2_eq_\_cong) + apply(rule eq_\_converter_reflI) + apply(rule WT_intro)+ + apply(rule parallel_converter2_eq_\_cong) + prefer 2 + apply(rule eq_\_converter_reflI) + apply(rule WT_intro)+ + apply(rule fact4) + prefer 3 + apply(subst attach_compose) + apply(fold converter_of_callee_id_oracle[where s="()"]) + apply(subst attach_parallel_fuse'[where f_init=id]) + apply(unfold converter_of_callee_id_oracle) + apply(subst eq_resource_on_UNIV_iff) + subgoal by (simp add: att_core[symmetric] att_rest[symmetric] real_s_core'_def real_s_rest'_def) + apply (rule WT_resource_of_oracle[where \="(\_full \\<^sub>\ \_full) \\<^sub>\ (((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full)) \\<^sub>\ (\_full \\<^sub>\ \_full))"]) + subgoal for s + apply (rule WT_calleeI) + apply (cases s) + apply(case_tac call) + apply(rename_tac [!] x) + apply(case_tac [!] x) + apply(rename_tac [!] y) + apply(case_tac [!] y) + apply(rename_tac [5-6] z) + apply (case_tac [5-6] z) + apply (auto simp add: fused_resource.fuse.simps parallel_eoracle_def) + done + apply simp + done + done +qed + + +subsection \A lazy construction and its DH reduction\ + + +subsubsection \Defining a lazy construction with an inlined sampler\ + +\ \"sample triple" and "basic core" states\ +type_synonym 'grp' st_state = "('grp' \ 'grp' \ 'grp') option" +type_synonym 'grp' bc_state = "('grp' st_state \ 'grp' cstate \ 'grp' cstate) \ 'grp' auth.state \ 'grp' auth.state" + +context + fixes sample_triple :: "('grp \ 'grp \ 'grp) spmf" +begin + +abbreviation basic_core_sinit :: "'grp bc_state" + where + "basic_core_sinit \ ((None, CState_Void, CState_Void), (auth.State_Void, {}), auth.State_Void, {})" + +fun basic_core_helper_base :: "('grp bc_state, unit, unit) oracle'" + where + "basic_core_helper_base ((s_key, CState_Void, s_cnv2), (auth.State_Void, parties1), s_auth2) _ = + (if auth.Alice \ parties1 + then return_spmf ((), (s_key, CState_Half 0, s_cnv2), (auth.State_Store \, parties1), s_auth2) + else return_pmf None)" + | "basic_core_helper_base _ _ = return_pmf None" + +definition basic_core_helper :: "('grp bc_state, icnv_alice + icnv_bob) handler" + where + "basic_core_helper \ (\state query. + let handle = \((sk, (sc1, sc2)), sa1, sa2). ((sk, (sc2, sc1)), sa2, sa1) in + let func = \h_s f s. map_spmf (h_s o snd) (f (h_s s) ()) in + let func_alc = func id basic_core_helper_base in + let func_bob = func handle basic_core_helper_base in + case_sum (\_. func_alc state) (\_. func_bob state) query)" + +fun basic_core_oracle_adv :: " unit + unit \ ('grp st_state \ 'grp auth.state, 'grp auth.iadv, 'grp auth.oadv) oracle'" + where + "basic_core_oracle_adv sel (None, auth.State_Store _, parties) (Inr (Inl _)) = do { + (gxy, gx, gy) \ sample_triple; + let out = case_sum (\_. gx) (\_. gy) sel; + return_spmf (Inr (Inl (auth.Out_Look out)), Some (gxy, gx, gy), auth.State_Store \, parties) + }" + | "basic_core_oracle_adv sel (Some dhs, auth.State_Store _, parties) (Inr (Inl _)) = + (case dhs of (gxy, gx, gy) \ + let out = case_sum (\_. gx) (\_. gy) sel in + return_spmf (Inr (Inl (auth.Out_Look out)), Some dhs, auth.State_Store \, parties))" + | "basic_core_oracle_adv _ (s_key, auth.State_Store _, parties) (Inr (Inr _)) = + return_spmf (Inr (Inr auth.Out_Fedit), s_key, auth.State_Collect \, parties)" + | "basic_core_oracle_adv _ _ _ = return_pmf None" + + +fun basic_core_oracle_usr_base :: "('grp bc_state, unit, 'grp) oracle'" + where + "basic_core_oracle_usr_base ((s_key, CState_Half _, s_cnv2), s_auth1, auth.State_Collect _, parties2) _ = + (let h_state = \k. ((Some k, CState_Full (0, \), s_cnv2), s_auth1, auth.State_Collected, parties2) in + if auth.Bob \ parties2 then + (case s_key of + None \ do { + (gxy, gx, gy) \ sample_triple; + return_spmf (gxy, h_state (gxy, gx, gy)) } + | Some (gxy, gx, gy) \ return_spmf (gxy, h_state (gxy, gx, gy))) + else return_pmf None)" + | "basic_core_oracle_usr_base ((Some dhs, CState_Full _, s_cnv2), s_auth1, auth.State_Collected, parties2) _ = + (case dhs of (gxy, gx, gy) \ + return_spmf (gxy, (Some dhs, CState_Full (0, \), s_cnv2), s_auth1, auth.State_Collected, parties2))" + | "basic_core_oracle_usr_base _ _ = return_pmf None" + +definition basic_core_oracle_usr :: "(_, key.iusr_alice + key.iusr_bob, _) oracle'" + where + "basic_core_oracle_usr \ (\state query. + let handle = \((sk, (sc1, sc2)), sa1, sa2). ((sk, (sc2, sc1)), sa2, sa1) in + let func = \h_o h_s f s. map_spmf (map_prod h_o h_s) (f (h_s s) ()) in + let func_alc = func (Inl o key.Out_Alice) id basic_core_oracle_usr_base in + let func_bob = func (Inr o key.Out_Bob) handle basic_core_oracle_usr_base in + case_sum (\_. func_alc state) (\_. func_bob state) query)" + +primcorec basic_core + where + "cpoke basic_core = (\(s_other, s_core) event. + map_spmf (Pair s_other) (parallel_handler auth.poke auth.poke s_core event))" + | "cfunc_adv basic_core = (\((s_key, s_cnv), s_auth1, s_auth2) iadv. + let handle = (\sel s_init h_out h_state query. + map_spmf + (\(out, (s_key', s_auth')). (h_out out, (s_key', s_cnv), h_state s_auth' s_auth1 s_auth2)) + (basic_core_oracle_adv sel (s_key, s_init) query)) in + case_sum (handle (Inl ()) s_auth1 Inl (\x y z. (x, z))) (handle (Inr ()) s_auth2 Inr (\x y z. (y, x))) iadv)" + | "cfunc_usr basic_core = + (let handle = map_sum (\_. Out_Activation_Alice) (\_. Out_Activation_Bob) in + basic_core_oracle_usr \\<^sub>O (\s q. map_spmf (Pair (handle q)) (basic_core_helper s q)))" + +primcorec lazy_core + where + "cpoke lazy_core = (\s. case_sum (\q. basic_core_helper s q) (cpoke basic_core s))" + | "cfunc_adv lazy_core = cfunc_adv basic_core " + | "cfunc_usr lazy_core = basic_core_oracle_usr" + +definition lazy_rest + where + "lazy_rest \ ideal_rest'" + +end + + +subsubsection \Defining a lazy construction with an external sampler\ + +context +begin + +private type_synonym ('grp', 'iadv_rest', 'iusr_rest') dh_inp = " + (('grp' auth.iadv + 'grp' auth.iadv) + 'iadv_rest') + (key.iusr_alice + key.iusr_bob) + (icnv_alice + icnv_bob) + 'iusr_rest'" + +private type_synonym ('grp', 'oadv_rest', 'ousr_rest') dh_out = " + (('grp' auth.oadv + 'grp' auth.oadv) + 'oadv_rest') + ('grp' key.ousr_alice + 'grp' key.ousr_bob) + (ocnv_alice + ocnv_bob) + 'ousr_rest'" + +fun interceptor_base_look :: " unit + unit \ 'grp st_state \ 'grp auth.state + \ ('grp auth.oadv_look \ 'grp st_state, unit, 'grp \ 'grp \ 'grp) gpv" + where + "interceptor_base_look sel (None, auth.State_Store _, parties) = do { + (gxy, gx, gy) \ Pause () Done; + let out = case_sum (\_. gx) (\_. gy) sel; + Done (auth.Out_Look out, Some (gxy, gx, gy)) }" + | "interceptor_base_look sel (Some dhs, auth.State_Store _, parties) = ( + case dhs of (gxy, gx, gy) \ + let out = case_sum (\_. gx) (\_. gy) sel in + Done (auth.Out_Look out, Some (gxy, gx, gy)))" + | "interceptor_base_look _ _ = Fail" + +fun interceptor_base_recv :: "'grp bc_state \ ('grp \ 'grp bc_state, unit, 'grp \ 'grp \ 'grp) gpv" + where + "interceptor_base_recv ((s_key, CState_Half _, s_cnv2), s_auth1, auth.State_Collect _, parties2) = ( + let h_state = \k. ((Some k, CState_Full (0, \), s_cnv2), s_auth1, auth.State_Collected, parties2) in + if auth.Bob \ parties2 then + case s_key of + None \ do { + (gxy, gx, gy) \ Pause () Done; + Done (gxy, h_state (gxy, gx, gy)) } + | Some (gxy, gx, gy) \ Done (gxy, h_state (gxy, gx, gy)) + else + Fail)" + | "interceptor_base_recv ((Some dhs, CState_Full _, s_cnv2), s_auth1, auth.State_Collected, parties2) = ( + case dhs of (gxy, gx, gy) \ + Done (gxy, (Some dhs, CState_Full (0, \), s_cnv2), s_auth1, auth.State_Collected, parties2))" + | "interceptor_base_recv _ = Fail" + +fun interceptor :: " _ \ (_, _, _) dh_inp \ (('grp, _, _) dh_out \ _, unit, 'grp \ 'grp \ 'grp) gpv" + where + "interceptor (sc, sr) (Inl (Inl (q))) = ( + let select_s = (case sc of ((sk, _), sa1, sa2) \ case_sum (\_. (sk, sa1)) (\_. (sk, sa2))) in + let handle_s = (\x. case sc of ((sk, (sc1, sc2)), sa1, sa2) \ ((x, (sc1, sc2)), sa1, sa2)) in + let func_look = (\sel h_o. do { + (o_look, dhs) \ interceptor_base_look (sel ()) (select_s (sel ())) ; + Done (Inl (Inl (h_o (Inr (Inl o_look)))), handle_s dhs, sr) }) in + let func_dfe = do { + (out, sc') \ lift_spmf (cfunc_adv (lazy_core undefined) sc q); + Done (Inl (Inl out), sc', sr) } in + case q of + (Inl (Inr (Inl _))) \ func_look Inl Inl + | (Inr (Inr (Inl _))) \ func_look Inr Inr + | _ \ func_dfe)" + | "interceptor (sc, sr) (Inl (Inr (q))) = do { + ((out, es), sr') \ lift_spmf (rfunc_adv lazy_rest sr q); + sc' \ lift_spmf (foldl_spmf (\a e. cpoke (lazy_core undefined) a e) (return_spmf sc) es); + Done (Inl (Inr out), (sc', sr')) }" + | "interceptor (sc, sr) (Inr (Inl (q))) = ( + let handle = \((sk, (sc1, sc2)), sa1, sa2). ((sk, (sc2, sc1)), sa2, sa1) in + let func_recv = (\h_o h_s. do { + (o_recv, sc') \ interceptor_base_recv (h_s sc); + Done (Inr (Inl (h_o o_recv)), h_s sc', sr) }) in + case_sum (\_. func_recv (Inl o key.Out_Alice) id) (\_. func_recv (Inr o key.Out_Bob) handle) q)" + | "interceptor (sc, sr) (Inr (Inr (q))) = do { + ((out, es), sr') \ lift_spmf (rfunc_usr lazy_rest sr q); + sc' \ lift_spmf (foldl_spmf (\a e. cpoke (lazy_core undefined) a e) (return_spmf sc) es); + Done (Inr (Inr out), (sc', sr')) }" + +end + + +subsubsection \Reduction to Diffie-Hellman game\ + +definition DH0_sample :: "('grp \ 'grp \ 'grp) spmf" + where + "DH0_sample = do { + x \ sample_uniform (order \); + y \ sample_uniform (order \); + return_spmf ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y) }" + +definition DH1_sample :: "('grp \ 'grp \ 'grp) spmf" + where + "DH1_sample = do { + x \ sample_uniform (order \); + y \ sample_uniform (order \); + z \ sample_uniform (order \); + return_spmf (\<^bold>g [^] z, \<^bold>g [^] x, \<^bold>g [^] y) }" + +lemma lossless_DH0_sample [simp]: "lossless_spmf DH0_sample" + by (auto simp add: DH0_sample_def sample_uniform_def intro: order_gt_0) + +lemma lossless_DH1_sample [simp]: "lossless_spmf DH1_sample" + by (auto simp add: DH1_sample_def sample_uniform_def intro: order_gt_0) + +definition DH_adversary_curry :: "('grp \ 'grp \ 'grp \ bool spmf) \ 'grp \ 'grp \ 'grp \ bool spmf" + where + "DH_adversary_curry \ (\A x y z. bind_spmf (return_spmf (z, x, y)) A)" + +definition DH_adversary + where + "DH_adversary D \ DH_adversary_curry (\xyz. + run_gpv (obsf_oracle (obsf_oracle (\(tpl, s) q. map_spmf (apsnd (Pair tpl) \ fst) (exec_gpv (\_ _. return_spmf (tpl, ())) (interceptor s q) ())))) + (obsf_distinguisher D) (OK (OK (xyz, basic_core_sinit, basic_rest_sinit))))" + +context +begin + +private abbreviation "S_ic_asm s_cnv1 s_cnv2 s_krn1 s_krn2 \ + let s_cnvs = {CState_Void} \ {CState_Half 0} \ {CState_Full (0, \)} in + let s_krns = {auth.State_Void} \ {auth.State_Store \} \ {auth.State_Collect \} \ {auth.State_Collected} in + s_cnv1 \ s_cnvs \ s_cnv2 \ s_cnvs \ s_krn1 \ s_krns \ s_krn2 \ s_krns" + +private inductive S_ic :: "('grp \ 'grp \ 'grp) spmf \ ('grp bc_state \ (unit \ unit) \ 'auth1_s_rest \ 'auth2_s_rest) spmf \ + (('grp \ 'grp \ 'grp) \ 'grp bc_state \ (unit \ unit) \ 'auth1_s_rest \ 'auth2_s_rest) spmf \ bool" + for \ :: "('grp \ 'grp \ 'grp) spmf" where + "S_ic \ (return_spmf (((None, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ((), ()), s_rest1, s_rest2)) + (map_spmf (\x. (x, (((None, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ((), ()), s_rest1, s_rest2))) \)" + if "S_ic_asm s_cnv1 s_cnv2 s_krn1 s_krn2" + | "S_ic \ (return_spmf (((Some x, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ((), ()), s_rest1, s_rest2)) + (return_spmf (x, (((Some x, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ((), ()), s_rest1, s_rest2)))" + if "S_ic_asm s_cnv1 s_cnv2 s_krn1 s_krn2" + +private lemma trace_eq_intercept: + defines "outs_adv \ ((UNIV <+> UNIV <+> UNIV) <+> UNIV <+> UNIV <+> UNIV) <+> UNIV <+> UNIV" + and "outs_usr \ (UNIV <+> UNIV) <+> (UNIV <+> UNIV) <+> UNIV <+> UNIV" + assumes "lossless_spmf sample_triple" + shows "trace_callee_eq (fused_resource.fuse (lazy_core sample_triple) lazy_rest) + (\(tpl, s) q. map_spmf (apsnd (Pair tpl) \ fst) (exec_gpv (\_ _. return_spmf (tpl, ())) (interceptor s q) ())) + (outs_adv <+> outs_usr) + (return_spmf (basic_core_sinit, basic_rest_sinit)) (pair_spmf sample_triple (return_spmf (basic_core_sinit, basic_rest_sinit)))" + (is "trace_callee_eq ?L ?R ?OI ?sl ?sr") +proof - + have auth_poke_alt[simplified split_def Let_def]: + "auth.poke = (\(sl, sr) q. let p = auth.case_event id q in + map_spmf (Pair sl) (if p \ sr then return_pmf None else return_spmf (insert p sr)))" + by (rule ext)+ (simp add: split_def Let_def split!: auth.event.splits) + + note S_ic_cases = S_ic.cases[where \=sample_triple, simplified] + note S_ic_intros = S_ic.intros[where \=sample_triple, simplified] + + note [simp] = assms(3)[unfolded lossless_spmf_def] mk_lossless_lossless[OF assms(3)] + fused_resource.fuse.simps lazy_rest_def basic_core_oracle_usr_def basic_core_helper_def + exec_gpv_bind spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def Let_def split_def + extend_state_oracle_def plus_eoracle_def parallel_eoracle_def map_fun_def + + have "trace_callee_eq ?L ?R ?OI ?sl ?sr" + unfolding assms(1,2) apply (rule trace'_eqI_sim_upto[where S="S_ic sample_triple"]) + subgoal Init_OK + by (simp add: map_spmf_conv_bind_spmf[symmetric] pair_spmf_alt_def S_ic_intros) + subgoal Out_OK for sl sr q + apply (cases q) + subgoal for q_adv + apply (cases q_adv) + subgoal for q_adv_core + apply (cases q_adv_core) + subgoal for q_acore1 + apply (cases q_acore1) + subgoal for q_drop by (erule S_ic_cases) simp_all + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (None, s_krn1, s_act1))" rule: interceptor_base_look.cases) auto + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (Some x, s_krn1, s_act1))" rule: interceptor_base_look.cases) auto + done + subgoal for q_fedit + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (None, s_krn1, s_act1), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) simp_all + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (Some x, s_krn1, s_act1), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) simp_all + done + done + done + subgoal for q_acore2 + apply (cases q_acore2) + subgoal for q_drop by (erule S_ic_cases) simp_all + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (None, s_krn2, s_act2))" rule: interceptor_base_look.cases) auto + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (Some x, s_krn2, s_act2))" rule: interceptor_base_look.cases) auto + done + subgoal for q_fedit + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (None, s_krn2, s_act2), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) simp_all + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (Some x, s_krn2, s_act2), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) simp_all + done + done + done + done + subgoal for q_adv_rest + apply (cases q_adv_rest) + subgoal for q_arest1 by (erule S_ic_cases) simp_all + subgoal for q_arest2 by (erule S_ic_cases) simp_all + done + done + subgoal for q_usr + apply (cases q_usr) + subgoal for q_usr_core + apply (cases q_usr_core) + subgoal for q_alice + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(((None, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_oracle_usr_base.cases) (auto split: option.splits) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(((Some x, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_oracle_usr_base.cases) (auto split: option.splits) + done + subgoal for q_bob + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(((None, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_oracle_usr_base.cases) (auto split: option.splits) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(((Some x, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_oracle_usr_base.cases) (auto split: option.splits) + done + done + subgoal for q_usr_rest + apply (cases q_usr_rest) + subgoal for q_act + apply (cases q_act) + subgoal for a_alice by (erule S_ic_cases) simp_all + subgoal for a_bob by (erule S_ic_cases) simp_all + done + subgoal for q_urest + apply (cases q_urest) + subgoal for q_urest1 by (erule S_ic_cases) simp_all + subgoal for q_urest2 by (erule S_ic_cases) simp_all + done + done + done + done + subgoal State_OK for sl sr q + apply (cases q) + subgoal for q_adv + apply (cases q_adv) + subgoal for q_adv_core + apply (cases q_adv_core) + subgoal for q_acore1 + apply (cases q_acore1) + subgoal for q_drop by (erule S_ic_cases) simp_all + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(Inl (), (None, s_krn1, s_act1))" rule: interceptor_base_look.cases) + apply (simp_all add: map_spmf_conv_bind_spmf[symmetric]) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def vimage_def intro!: trace_eq_simcl_map S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(Inl (), (Some x, s_krn1, s_act1))" rule: interceptor_base_look.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + subgoal for q_fedit + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (None, s_krn1, s_act1), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) + (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inl (), (Some x, s_krn1, s_act1), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) + (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + done + done + subgoal for q_acore2 + apply (cases q_acore2) + subgoal for q_drop by (erule S_ic_cases) simp_all + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(Inr (), (None, s_krn2, s_act2))" rule: interceptor_base_look.cases) + apply (simp_all add: map_spmf_conv_bind_spmf[symmetric]) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def vimage_def intro!: trace_eq_simcl_map S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(Inr (), (Some x, s_krn2, s_act2))" rule: interceptor_base_look.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + subgoal for q_fedit + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (None, s_krn2, s_act2), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) + (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + by (simp, case_tac "(Inr (), (Some x, s_krn2, s_act2), Inr (Inr q_fedit))" rule: basic_core_oracle_adv.cases) + (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + done + done + done + subgoal for q_adv_rest + apply (cases q_adv_rest) + subgoal for q_urest1 + apply (erule S_ic_cases) + subgoal + apply clarsimp + apply (subst bind_commute_spmf) + apply (subst (2) bind_commute_spmf) + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, (_, x), _). x" + and p="\(a, (b, _), d). (a, b, d)" and f="\(a, b, d) c. (a, (b, c), d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal + apply clarsimp + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, (_, x), _). x" + and p="\(a, (b, _), d). (a, b, d)" and f="\(a, b, d) c. (a, (b, c), d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + subgoal for q_urest2 + apply (erule S_ic_cases) + subgoal + apply clarsimp + apply (subst bind_commute_spmf) + apply (subst (2) bind_commute_spmf) + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, _, _, x). x" + and p="\(a, b, c, _). (a, b, c)" and f="\(a, b, c) d. (a, b, c, d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal + apply clarsimp + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, _, _, x). x" + and p="\(a, b, c, _). (a, b, c)" and f="\(a, b, c) d. (a, b, c, d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + done + done + subgoal for q_usr + apply (cases q_usr) + subgoal for q_usr_core + apply (cases q_usr_core) + subgoal for q_alice + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((None, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_oracle_usr_base.cases) + proof (goal_cases) + case (1 s_key _ s_cnv2 s_auth1 _ parties2 _) + then show ?case by (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def vimage_def intro!: trace_eq_simcl_map S_ic_intros) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((Some x, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_oracle_usr_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + subgoal for q_bob + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((None, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_oracle_usr_base.cases) + proof (goal_cases) + case (1 s_key _ s_cnv2 s_auth1 _ parties2 _) + then show ?case by (auto simp add: map_spmf_conv_bind_spmf[symmetric] cond_spmf_fst_def vimage_def intro!: trace_eq_simcl_map S_ic_intros) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((Some x, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_oracle_usr_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + done + subgoal for q_usr_rest + apply (cases q_usr_rest) + subgoal for q_act + apply (cases q_act) + subgoal for a_alice + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((None, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_helper_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((Some x, s_cnv1, s_cnv2), (s_krn1, s_act1), s_krn2, s_act2), ())" rule: basic_core_helper_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + subgoal for a_bob + apply (erule S_ic_cases) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((None, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_helper_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal for s_cnv1 s_cnv2 s_krn1 s_krn2 x s_act1 s_act2 s_rest1 s_rest2 + apply (simp, case_tac "(((Some x, s_cnv2, s_cnv1), (s_krn2, s_act2), s_krn1, s_act1), ())" rule: basic_core_helper_base.cases) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + done + done + subgoal for q_urest + apply (cases q_urest) + subgoal for q_urest1 + apply (erule S_ic_cases) + subgoal + apply clarsimp + apply (subst bind_commute_spmf) + apply (subst (2) bind_commute_spmf) + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, (_, x), _). x" + and p="\(a, (b, _), d). (a, b, d)" and f="\(a, b, d) c. (a, (b, c), d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal + apply clarsimp + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, (_, x), _). x" + and p="\(a, (b, _), d). (a, b, d)" and f="\(a, b, d) c. (a, (b, c), d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + subgoal for q_urest2 + apply (erule S_ic_cases) + subgoal + apply clarsimp + apply (subst bind_commute_spmf) + apply (subst (2) bind_commute_spmf) + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, _, _, x). x" + and p="\(a, b, c, _). (a, b, c)" and f="\(a, b, c) d. (a, b, c, d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_ic_intros) + subgoal + apply clarsimp + apply (subst (1 2) cond_spmf_fst_bind) + apply (subst (1 2) cond_spmf_fst_bind) + apply (clarsimp intro!: trace_eq_simcl_bind simp add: auth_poke_alt set_scale_spmf split: if_split_asm) + apply (subst (asm) (1 2 3 4) foldl_spmf_helper2[where acc="return_spmf _" and q="\(_, _, _, x). x" + and p="\(a, b, c, _). (a, b, c)" and f="\(a, b, c) d. (a, b, c, d)", simplified]) + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_ic_intros) + done + done + done + done + done + done + + then show ?thesis by simp +qed + +private abbreviation "dummy x \ TRY map_spmf OK x ELSE return_spmf Fault" + +lemma reduction: "advantage D (obsf_resource (RES (fused_resource.fuse (lazy_core DH1_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit))) + (obsf_resource (RES (fused_resource.fuse (lazy_core DH0_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit))) = ddh.advantage \ (DH_adversary D)" +proof - + have fact1: "bind_spmf (DH0_sample) A = do { + x \ sample_uniform (order \); + y \ sample_uniform (order \); + (DH_adversary_curry A) (\<^bold>g [^] x) (\<^bold>g [^] y) (\<^bold>g [^] (x * y)) + }" for A by (simp add: DH0_sample_def DH_adversary_curry_def nat_pow_pow) + + have fact2: "bind_spmf DH1_sample A = do { + x \ sample_uniform (order \); + y \ sample_uniform (order \); + z \ sample_uniform (order \); + (DH_adversary_curry A) (\<^bold>g [^] x) (\<^bold>g [^] y) (\<^bold>g [^] (z)) + }" for A by (simp add: DH1_sample_def DH_adversary_curry_def) + + { + fix sample_triple :: "('grp \ 'grp \ 'grp) spmf" + assume *: "lossless_spmf sample_triple" + define s_init where "s_init \ (basic_core_sinit, basic_rest_sinit)" + have [unfolded s_init_def, simp]: "dummy (dummy (return_spmf s_init)) = return_spmf (OK (OK s_init))" by auto + have [unfolded s_init_def, simp]: "dummy (dummy (pair_spmf sample_triple (return_spmf s_init))) = + map_spmf (OK \ OK) (pair_spmf sample_triple (return_spmf s_init))" + using * by (simp add: o_def map_spmf_conv_bind_spmf pair_spmf_alt_def) + + have "connect D (RES (obsf_oracle (obsf_oracle (fused_resource.fuse (lazy_core sample_triple) lazy_rest))) (OK (OK (basic_core_sinit, basic_rest_sinit)))) = + bind_spmf (map_spmf (OK o OK) (pair_spmf sample_triple (return_spmf (basic_core_sinit, basic_rest_sinit)))) + (run_gpv (obsf_oracle (obsf_oracle (\(tpl, s) q. map_spmf ((apsnd (Pair tpl)) o fst) (exec_gpv (\_ _. return_spmf (tpl, ())) (interceptor s q) ())))) D)" for D + apply (simp add: connect_def exec_gpv_resource_of_oracle spmf.map_comp) + apply (subst return_bind_spmf[where x="OK (OK (basic_core_sinit, basic_rest_sinit))", symmetric]) + apply (rule trace_callee_eq_run_gpv[where ?I1.0="(\_. True)" and ?I2.0="(\_. True)" and \=\_full and A=UNIV]) + subgoal by (rule trace_eq_intercept[OF *, THEN trace_callee_eq_obsf_oracleI, THEN trace_callee_eq_obsf_oracleI, simplified]) + by (simp_all add: * pair_spmf_alt_def) + } note detach_sampler = this + + show ?thesis + unfolding advantage_def + apply (subst (1 2) spmf_obsf_distinguisher_obsf_resource_True[symmetric]) + apply (subst (1 2) obsf_resource_of_oracle)+ + by (simp add: detach_sampler pair_spmf_alt_def bind_map_spmf fact1 fact2) + (simp add: ddh.advantage_def ddh.ddh_0_def ddh.ddh_1_def DH_adversary_def) +qed + +end + + +subsection \Proving the trace-equivalence of simplified Ideal and Lazy constructions\ + +context +begin + +private abbreviation "isample_nat \ sample_uniform (order \)" +private abbreviation "isample_key \ spmf_of_set (carrier \)" +private abbreviation "isample_pair_nn \ pair_spmf isample_nat isample_nat" +private abbreviation "isample_pair_nk \ pair_spmf isample_nat isample_key" + +private inductive S_il :: "(('grp astate \ unit) \ estate \ 'grp key.state) spmf \ 'grp bc_state spmf \ bool" + where +\ \(Auth1 =a)@(Auth2 =0)\ + sil_0_0: "S_il (return_spmf ((None, ()), (False, (EState_Void, s_act1), EState_Void, s_act2), key.PState_Store, {})) + (return_spmf ((None, CState_Void, CState_Void), (auth.State_Void, s_act1), auth.State_Void, s_act2))" +\ \../(Auth1 =a)@(Auth2 =0) \# wl\ + | sil_1_0: "S_il (return_spmf ((None, ()), (False, (EState_Store, s_act1), EState_Void, s_act2), key.PState_Store, {})) + (return_spmf ((None, CState_Half 0, CState_Void), (auth.State_Store \, s_act1), auth.State_Void, s_act2))" + if "auth.Alice \ s_act1" + | sil_2_0: "S_il (map_spmf (\k. ((None, ()), (True, (EState_Collect, s_act1), EState_Void, s_act2), key.State_Store k, {})) isample_key) + (return_spmf ((None, CState_Half 0, CState_Void), (auth.State_Collect \, s_act1), auth.State_Void, s_act2))" + if "auth.Alice \ s_act1" +\ \../(Auth1 =a)@(Auth2 =0) \# look\ + | sil_1'_0: "S_il (map_spmf (\y. ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (False, (EState_Store, s_act1), EState_Void, s_act2), key.PState_Store, {})) isample_nat) + (map_spmf (\yz. ((Some (\<^bold>g [^] snd yz, \<^bold>g [^] (x :: nat), \<^bold>g [^] fst yz), CState_Half 0, CState_Void), (auth.State_Store \, s_act1), auth.State_Void, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" + | sil_2'_0: "S_il (map_spmf (\yk. ((Some (\<^bold>g [^] x, \<^bold>g [^] fst yk), ()), (True, (EState_Collect, s_act1), EState_Void, s_act2), key.State_Store (snd yk), {})) isample_pair_nk) + (map_spmf (\yz. ((Some (\<^bold>g [^] snd yz, \<^bold>g [^] (x :: nat), \<^bold>g [^] fst yz), CState_Half 0, CState_Void), (auth.State_Collect \, s_act1), auth.State_Void, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" +\ \(Auth1 =a)@(Auth2 =1)\ + | sil_0_1: "S_il (return_spmf ((None, ()), (False, (EState_Void, s_act1), EState_Store, s_act2), key.PState_Store, {})) + (return_spmf ((None, CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Store \, s_act2))" + if "auth.Alice \ s_act2" +\ \../(Auth1 =a)@(Auth2 =1) \# wl\ + | sil_1_1: "S_il (return_spmf ((None, ()), (False, (EState_Store, s_act1), EState_Store, s_act2), key.PState_Store, {})) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Store \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" + | sil_2_1: "S_il (map_spmf (\k. ((None, ()), (True, (EState_Collect, s_act1), EState_Store, s_act2), key.State_Store k, s_actk)) isample_key) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Store \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" + | sil_3_1: "S_il (return_spmf ((None, ()), (True, (EState_Collect, s_act1), EState_Store, s_act2), key.State_Store k, s_actk)) + (map_spmf (\xy. ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] fst xy, \<^bold>g [^] snd xy), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Store \, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \../(Auth1 =a)@(Auth2 =1) \# look\ + | sil_1c_1c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (False, (EState_Store, s_act1), EState_Store, s_act2), key.PState_Store, {})) + (map_spmf (\z. ((Some (\<^bold>g [^] z, \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Store \, s_act2)) isample_nat)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" + | sil_2c_1c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Store, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] z, \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Store \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" and "k = \<^bold>g [^] z" and "z \ set_spmf isample_nat" + | sil_3c_1c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Store, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Store \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \(Auth1 =a)@(Auth2 =2)\ + | sil_0_2: "S_il (map_spmf (\k. ((None, ()), (True, (EState_Void, s_act1), EState_Collect, s_act2), key.State_Store k, {})) isample_key) + (return_spmf ((None, CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act2" +\ \../(Auth1 =a)@(Auth2 =2) \# wl\ + | sil_1_2: "S_il (map_spmf (\k. ((None, ()), (True, (EState_Store, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) isample_key) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "key.Bob \ s_actk" + | sil_2_2: "S_il (map_spmf (\k. ((None, ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) isample_key) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" + | sil_3_2: "S_il (return_spmf ((None, ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (map_spmf (\xy. ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] fst xy, \<^bold>g [^] snd xy), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collect \, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \../(Auth1 =a)@(Auth2 =2) \# look\ + | sil_1c_2c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Store, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] z, \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" and "z \ set_spmf isample_nat" + | sil_2c_2c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] z, \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" and "k = \<^bold>g [^] z" and "z \ set_spmf isample_nat" + | sil_3c_2c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collect \, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2 \ key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \(Auth1 =a)@(Auth2 =3)\ +\ \../(Auth1 =a)@(Auth2 =3) \# wl\ + | sil_1_3: "S_il (return_spmf ((None, ()), (True, (EState_Store, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (map_spmf (\xy. ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] fst xy, \<^bold>g [^] snd xy), CState_Full (0, \), CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collected, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" + | sil_2_3: "S_il (return_spmf ((None, ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (map_spmf (\xy. ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] fst xy, \<^bold>g [^] snd xy), CState_Full (0, \), CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collected, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" and "k = \<^bold>g [^] z" + | sil_3_3: "S_il (return_spmf ((None, ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (map_spmf (\xy. ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] fst xy, \<^bold>g [^] snd xy), CState_Full (0, \), CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collected, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \../(Auth1 =a)@(Auth2 =3) \# look\ + | sil_1c_3c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Store, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Full (0, \), CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collected, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" + | sil_2c_3c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Full (0, \), CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collected, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1 \ key.Bob \ s_actk" and "k = \<^bold>g [^] z" + | sil_3c_3c: "S_il (return_spmf ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (True, (EState_Collect, s_act1), EState_Collect, s_act2), key.State_Store k, s_actk)) + (return_spmf ((Some (\<^bold>g [^] (z :: nat), \<^bold>g [^] (x :: nat), \<^bold>g [^] (y :: nat)), CState_Full (0, \), CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collected, s_act2))" + if "auth.Alice \ s_act1" and "auth.Alice \ s_act2" and "auth.Bob \ s_act2" and "key.Alice \ s_actk" and "auth.Bob \ s_act1" and "key.Bob \ s_actk" and "k = \<^bold>g [^] z" +\ \(Auth1 =a)@(Auth2 =1')\ + | sil_0_1': "S_il (map_spmf (\x. ((Some (\<^bold>g [^] x, \<^bold>g [^] y), ()), (False, (EState_Void, s_act1), EState_Store, s_act2), key.PState_Store, {})) isample_nat) + (map_spmf (\xz. ((Some (\<^bold>g [^] snd xz, \<^bold>g [^] fst xz, \<^bold>g [^] (y :: nat)), CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Store \, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act2" +\ \(Auth1 =a)@(Auth2 =2')\ + | sil_0_2': "S_il (map_spmf (\xk. ((Some (\<^bold>g [^] fst xk, \<^bold>g [^] y), ()), (True, (EState_Void, s_act1), EState_Collect, s_act2), key.State_Store (snd xk), {})) isample_pair_nk) + (map_spmf (\xz. ((Some (\<^bold>g [^] snd xz, \<^bold>g [^] fst xz, \<^bold>g [^] (y :: nat)), CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Collect \, s_act2)) isample_pair_nn)" + if "auth.Alice \ s_act2" + +private lemma trac_eq_core_il: "trace_core_eq ideal_core' (lazy_core DH1_sample) + ((UNIV <+> UNIV) <+> UNIV <+> UNIV) ((UNIV <+> UNIV <+> UNIV) <+> UNIV <+> UNIV <+> UNIV) (UNIV <+> UNIV) + (return_spmf ideal_s_core') (return_spmf basic_core_sinit)" +proof - + have isample_key_conv_nat[simplified map_spmf_conv_bind_spmf]: + "map_spmf f isample_key = map_spmf (\x. f (\<^bold>g [^] x)) isample_nat" for f + unfolding sample_uniform_def carrier_conv_generator + by (simp add: map_spmf_of_set_inj_on[OF inj_on_generator, symmetric] spmf.map_comp o_def) + + have [simp]: "weight_spmf isample_nat = 1" + by (simp add: finite_carrier order_gt_0_iff_finite) + + have [simp]: "weight_spmf isample_key = 1" + by (simp add: carrier_not_empty cyclic_group.finite_carrier cyclic_group_axioms) + + have [simp]: "mk_lossless isample_nat = isample_nat" + by (simp add: mk_lossless_def) + + have [simp]: "mk_lossless isample_pair_nn = isample_pair_nn" + by (simp add: mk_lossless_def) + + have [simp]: "mk_lossless isample_pair_nk = isample_pair_nk" + by (simp add: mk_lossless_def) + + note [simp] = basic_core_helper_def basic_core_oracle_usr_def eleak_def DH1_sample_def + Let_def split_def exec_gpv_bind spmf.map_comp o_def map_bind_spmf bind_map_spmf bind_spmf_const + + show ?thesis + apply (rule trace_core_eq_simI_upto[where S=S_il]) + subgoal Init_OK + by (simp add: ideal_s_core'_def einit_def sil_0_0) + subgoal POut_OK for sl sr query + apply (cases query) + subgoal for e_usrs + apply (cases e_usrs) + subgoal for e_alice by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + subgoal for e_bob by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + done + subgoal for e_chns + apply (cases e_chns) + subgoal for e_chn1 + apply (cases e_chn1) + subgoal for e_shell + apply (cases e_shell) + subgoal a_alice by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + subgoal a_bob by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + done + done + subgoal for e_chn2 + apply (cases e_chn2) + subgoal for e_shell + apply (cases e_shell) + subgoal a_alice by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + subgoal a_bob by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + done + done + done + done + subgoal PState_OK for sl sr query + apply (cases query) + subgoal for e_usrs + apply (cases e_usrs) + subgoal for e_alice + proof (erule S_il.cases, goal_cases) + case (26 s_act2 y s_act1) \ \Corresponds to @{thm [source] sil_0_1'}\ + then show ?case + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (simp add: pair_spmf_alt_def map_spmf_conv_bind_spmf) + apply (rule trace_eq_simcl_bindI) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros) + next + case (27 s_act2 y s_act1) \ \Corresponds to @{thm [source] sil_0_2'}\ + then show ?case + apply (clarsimp simp add: pair_spmf_alt_def isample_key_conv_nat) + apply (simp add: bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_il.intros trace_eq_simcl_map) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + subgoal for e_bob + proof (erule S_il.cases, goal_cases) + case (4 s_act1 x s_act2) \ \Corresponds to @{thm [source] sil_1'_0}\ + then show ?case + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (simp add: pair_spmf_alt_def map_spmf_conv_bind_spmf) + apply (rule trace_eq_simcl_bindI) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros) + next + case (5 s_act1 x s_act2) \ \Corresponds to @{thm [source] sil_2'_0}\ + then show ?case + apply (clarsimp simp add: pair_spmf_alt_def isample_key_conv_nat) + apply (simp add: bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: S_il.intros trace_eq_simcl_map) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + done + subgoal for e_chns + apply (cases e_chns) + subgoal for e_auth1 + apply (cases e_auth1) + subgoal for e_shell + apply (cases e_shell) + subgoal a_alice by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + subgoal a_bob by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + done + done + subgoal for e_auth2 + apply (cases e_auth2) + subgoal for e_shell + apply (cases e_shell) + subgoal a_alice by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + subgoal a_bob by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + done + done + done + done + subgoal AOut_OK for sl sr query + apply (cases query) + subgoal for q_auth1 + apply (cases q_auth1) + subgoal for q_drop by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look by (erule S_il.cases) (simp_all del: bind_spmf_const add: pair_spmf_alt_def, clarsimp+) + subgoal for q_fedit by (erule S_il.cases) (simp_all del: bind_spmf_const add: pair_spmf_alt_def, clarsimp+) + done + done + subgoal for q_auth2 + apply (cases q_auth2) + subgoal for q_drop by (erule S_il.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look by (erule S_il.cases) (simp_all del: bind_spmf_const add: pair_spmf_alt_def, clarsimp+) + subgoal for q_fedit by (erule S_il.cases) (simp_all del: bind_spmf_const add: pair_spmf_alt_def, clarsimp+) + done + done + done + subgoal AState_OK for sl sr query + apply (cases query) + subgoal for q_auth1 + apply (cases q_auth1) + subgoal for q_drop by (erule S_il.cases) auto + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + proof (erule S_il.cases, goal_cases) + case (2 s_act1 s_act2) \ \Corresponds to @{thm [source] sil_1_0}\ + then show ?case + apply simp + apply (subst (1 2 3) bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + by (auto intro: trace_eq_simcl_bindI S_il.intros) + next + case (7 s_act1 s_act2) \ \Corresponds to @{thm [source] sil_1_1}\ + then show ?case + apply simp + apply (subst (1 2 3) bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf pair_spmf_alt_def isample_key_conv_nat) + apply (rule trace_eq_simcl_bindI) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros) + next + case (14 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_1_2}\ + then show ?case + apply clarsimp + apply (subst bind_commute_spmf, subst (2) bind_commute_spmf) + apply (subst (1 2 3 4) bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf pair_spmf_alt_def isample_key_conv_nat) + apply (subst (1 2) bind_bind_conv_pair_spmf) + by (auto intro!: trace_eq_simcl_bindI S_il.intros) + next + case (20 s_act1 s_act2 s_actk k z) \ \Corresponds to @{thm [source] sil_1_3}\ + then show ?case + apply simp + apply (subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf) + by (auto intro!: trace_eq_simcl_bindI S_il.intros) + qed (auto simp add: map_spmf_conv_bind_spmf, + auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + subgoal for q_fedit + proof (erule S_il.cases, goal_cases) + case (4 s_act1 x s_act2) \ \Corresponds to @{thm [source] sil_1'_0}\ + then show ?case + apply simp + apply (subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + by (auto intro: S_il.intros trace_eq_simcl.base) + next + case (10 s_act1 s_act2 x y) \ \Corresponds to @{thm [source] sil_1c_1c}\ + then show ?case + apply (clarsimp simp add: pair_spmf_alt_def isample_key_conv_nat) + apply (simp add: map_spmf_conv_bind_spmf[symmetric]) + by (auto intro!: trace_eq_simcl_map S_il.intros) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric], + auto intro: S_il.intros trace_eq_simcl.base trace_eq_simcl_map) + done + done + subgoal for q_auth2 + apply (cases q_auth2) + subgoal for q_drop by (erule S_il.cases) auto + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + proof (erule S_il.cases, goal_cases) + case (6 s_act2 s_act1) \ \Corresponds to @{thm [source] sil_0_1}\ + then show ?case + apply clarsimp + apply (subst (1 2) bind_commute_spmf) + apply (subst (1 3) bind_bind_conv_pair_spmf) + apply (subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + by (auto intro: trace_eq_simcl_bindI S_il.intros) + next + case (7 s_act1 s_act2) \ \Corresponds to @{thm [source] sil_1_1}\ + then show ?case + apply clarsimp + apply (subst (1 2) bind_commute_spmf) + apply (subst (1 3) bind_bind_conv_pair_spmf) + apply (subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf) + apply (subst pair_spmf_alt_def) + apply (subst bind_spmf_assoc) + apply (rule trace_eq_simcl_bindI) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros) + next + case (8 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_2_1}\ + then show ?case + apply clarsimp + apply (subst (2) bind_commute_spmf, subst (1 3) bind_commute_spmf) + apply (subst (2) bind_commute_spmf) + apply (subst (2 4) bind_bind_conv_pair_spmf) + apply (clarsimp simp add: bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf) + apply (simp add: pair_spmf_alt_def isample_key_conv_nat) + apply (subst (1 2) bind_bind_conv_pair_spmf) + by (auto intro!: trace_eq_simcl_bindI S_il.intros) + next + case (9 s_act1 s_act2 s_actk k z) \ \Corresponds to @{thm [source] sil_3_1}\ + then show ?case + apply (clarsimp simp del: bind_spmf_const simp add: pair_spmf_alt_def) + apply (subst (1 2) bind_commute_spmf) + apply (subst (1 2) bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + apply (simp_all add: inj_on_def map_spmf_conv_bind_spmf) + by (auto intro: trace_eq_simcl_bindI S_il.intros) + qed (auto simp del: bind_spmf_const simp add: map_spmf_conv_bind_spmf, + auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: S_il.intros trace_eq_simcl.base) + subgoal for q_fedit + proof (erule S_il.cases, goal_cases) + case (10 s_act1 s_act2 x y) \ \Corresponds to @{thm [source] sil_1c_1c}\ + then show ?case + apply simp + apply (clarsimp simp add: map_spmf_conv_bind_spmf pair_spmf_alt_def isample_key_conv_nat) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + by (auto intro!: S_il.intros trace_eq_simcl_map) + next + case (26 s_act2 y s_act1) \ \Corresponds to @{thm [source] sil_0_1'}\ + then show ?case + apply simp + apply (subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + by (auto intro: S_il.intros trace_eq_simcl.base) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric], + auto intro: S_il.intros trace_eq_simcl.base trace_eq_simcl_map) + done + done + done + subgoal UOut_OK for sl sr query + apply (cases query) + subgoal for q_alice + apply (erule S_il.cases) + by (auto simp add: pair_spmf_alt_def isample_key_conv_nat) + subgoal for q_bob + apply (erule S_il.cases) + by (auto simp add: pair_spmf_alt_def isample_key_conv_nat) + done + subgoal UState_OK for sl sr query + apply (cases query) + subgoal for q_alice + proof (erule S_il.cases, goal_cases) + case (14 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_1_2}\ + then show ?case + apply (clarsimp) + apply (subst (2) bind_commute_spmf, subst bind_commute_spmf) + apply (subst bind_bind_conv_pair_spmf, subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply(subst (1 2) inv_into_f_f) + by (auto simp add: inj_on_def intro: S_il.intros trace_eq_simcl.base) + next + case (15 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_2_2}\ + then show ?case + apply (clarsimp) + apply (subst (2) bind_commute_spmf, subst bind_commute_spmf) + apply (subst bind_bind_conv_pair_spmf, subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply(subst (1 2) inv_into_f_f) + by (auto simp add: inj_on_def intro: S_il.intros trace_eq_simcl.base) + qed (auto simp add: map_spmf_conv_bind_spmf[symmetric], + auto intro: S_il.intros trace_eq_simcl.base trace_eq_simcl_map) + subgoal for q_bob + proof (erule S_il.cases, goal_cases) + case (8 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_2_1}\ + then show ?case + apply clarsimp + apply (subst (2) bind_commute_spmf, subst bind_commute_spmf) + apply (subst (2) bind_bind_conv_pair_spmf, subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply (subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + by (auto simp add: inj_on_def intro: S_il.intros trace_eq_simcl.base) + next + case (15 s_act1 s_act2 s_actk) \ \Corresponds to @{thm [source] sil_2_2}\ + then show ?case + apply clarsimp + apply (subst (2) bind_commute_spmf, subst bind_commute_spmf) + apply (subst (2) bind_bind_conv_pair_spmf, subst bind_bind_conv_pair_spmf) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply (subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + apply (subst (1 2) inv_into_f_f) + by (auto simp add: inj_on_def intro: S_il.intros trace_eq_simcl.base) + qed(auto simp add: map_spmf_conv_bind_spmf[symmetric], + auto intro: S_il.intros trace_eq_simcl.base trace_eq_simcl_map) + done + done +qed + +lemma connect_ideal: "connect D (obsf_resource ideal_resource) = + connect D (obsf_resource (RES (fused_resource.fuse (lazy_core DH1_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit)))" +proof - + have fact1: "trace_rest_eq ideal_rest' ideal_rest' UNIV UNIV s s" for s + by (rule rel_rest'_into_trace_rest_eq[where S="(=)" and M="(=)"]) (simp_all add: eq_onp_def rel_rest'_eq) + + have fact2: "\_full \\<^sub>\ \_full \c callee_of_rest ideal_rest' s \" for s + by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + + have fact3: "\_full \\<^sub>\ (\_full \\<^sub>\ \_full) \c callee_of_core ideal_core' s \" for s + by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + + have fact4: "\_full \\<^sub>\ (\_full \\<^sub>\ \_full) \c callee_of_core (lazy_core xyz) s \" for xyz s + by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + + show ?thesis + apply (rule connect_cong_trace[where A="UNIV" and \=\_full]) + apply (rule trace_eq_obsf_resourceI) + subgoal + apply (simp add: attach_ideal) + apply (rule fuse_trace_eq[where \E=\_full and \CA=\_full and \CU=\_full and \RA=\_full and \RU=\_full, simplified]) + by (simp_all add: ideal_s_rest'_def lazy_rest_def trac_eq_core_il[simplified] fact1 fact2 fact3 fact4) + by (simp_all add: attach_ideal) +qed + + +end + + +subsection \Proving the trace-equivalence of simplified Real and Lazy constructions\ + +context +begin + +private abbreviation "rsample_nat \ sample_uniform (order \)" +private abbreviation "rsample_pair_nn \ pair_spmf rsample_nat rsample_nat" + +private inductive S_rl :: "((unit \ 'grp cstate \ 'grp cstate) \ 'grp auth.state \ 'grp auth.state) spmf + \ (('grp st_state \ 'grp cstate \ 'grp cstate) \ 'grp auth.state \ 'grp auth.state) spmf \ bool" + where +\ \(Auth1 =a)@(Auth2 =0)\ + srl_0_0: "S_rl (return_spmf (((), CState_Void, CState_Void), (auth.State_Void, s_act1), auth.State_Void, s_act2)) + (return_spmf ((None, CState_Void, CState_Void), (auth.State_Void, s_act1), (auth.State_Void, s_act2)))" +\ \../(Auth1 =a)@(Auth2 =0) \# wl\ + | srl_1_0: "S_rl (map_spmf (\x. (((), CState_Half x, CState_Void), (auth.State_Store (\<^bold>g [^] x), s_act1), auth.State_Void, s_act2)) rsample_nat) + (return_spmf ((None, CState_Half 0, CState_Void), (auth.State_Store \, s_act1), auth.State_Void, s_act2))" + | srl_2_0: "S_rl (map_spmf (\x. (((), CState_Half x, CState_Void), (auth.State_Collect (\<^bold>g [^] x), s_act1), auth.State_Void, s_act2)) rsample_nat) + (return_spmf ((None, CState_Half 0, CState_Void), (auth.State_Collect \, s_act1), auth.State_Void, s_act2))" +\ \../(Auth1 =a)@(Auth2 =0) \# look\ + | srl_1'_0: "S_rl (return_spmf (((), CState_Half x, CState_Void), (auth.State_Store (\<^bold>g [^] x), s_act1), auth.State_Void, s_act2)) + (map_spmf (\y. ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Void), (auth.State_Store \, s_act1), auth.State_Void, s_act2)) rsample_nat)" + | srl_2'_0: "S_rl (return_spmf (((), CState_Half x, CState_Void), (auth.State_Collect (\<^bold>g [^] x), s_act1), auth.State_Void, s_act2)) + (map_spmf (\y. ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Void), (auth.State_Collect \, s_act1), auth.State_Void, s_act2)) rsample_nat)" +\ \(Auth1 =a)@(Auth2 =1)\ + | srl_0_1: "S_rl (map_spmf (\y. (((), CState_Void, CState_Half y), (auth.State_Void, s_act1), auth.State_Store (\<^bold>g [^] y), s_act2)) rsample_nat) + (return_spmf ((None, CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Store \, s_act2))" +\ \../(Auth1 =a)@(Auth2 =1) \# wl\ + | srl_1_1: "S_rl (map_spmf (\yx. (((), CState_Half (snd yx), CState_Half (fst yx)), (auth.State_Store (\<^bold>g [^] snd yx), s_act1), auth.State_Store (\<^bold>g [^] fst yx), s_act2)) rsample_pair_nn) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Store \, s_act2))" + | srl_2_1: "S_rl (map_spmf (\yx. (((), CState_Half (snd yx), CState_Half (fst yx)), (auth.State_Collect (\<^bold>g [^] snd yx), s_act1), auth.State_Store (\<^bold>g [^] fst yx), s_act2)) rsample_pair_nn) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Store \, s_act2))" +\ \../(Auth1 =a)@(Auth2 =1) \# look\ + | srl_1c_1c: "S_rl (return_spmf (((), CState_Half x, CState_Half y), (auth.State_Store (\<^bold>g [^] x), s_act1), auth.State_Store (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Store \, s_act2))" + | srl_2c_1c: "S_rl (return_spmf (((), CState_Half x, CState_Half y), (auth.State_Collect (\<^bold>g [^] x), s_act1), auth.State_Store (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Store \, s_act2))" + | srl_3c_1c: "S_rl (return_spmf (((), CState_Half x, CState_Full (y, z)), (auth.State_Collected, s_act1), auth.State_Store (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some (z, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Store \, s_act2))" + if "z = (\<^bold>g [^] x) [^] y" +\ \(Auth1 =a)@(Auth2 =2)\ + | srl_0_2: "S_rl (map_spmf (\y. (((), CState_Void, CState_Half y), (auth.State_Void, s_act1), auth.State_Collect (\<^bold>g [^] y), s_act2)) rsample_nat) + (return_spmf ((None, CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Collect \, s_act2))" +\ \../(Auth1 =a)@(Auth2 =2) \# wl\ + | srl_1_2: "S_rl (map_spmf (\yx. (((), CState_Half (snd yx), CState_Half (fst yx)), (auth.State_Store (\<^bold>g [^] snd yx), s_act1), auth.State_Collect (\<^bold>g [^] fst yx), s_act2)) rsample_pair_nn) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collect \, s_act2))" + | srl_2_2: "S_rl (map_spmf (\yx. (((), CState_Half (snd yx), CState_Half (fst yx)), (auth.State_Collect (\<^bold>g [^] snd yx), s_act1), auth.State_Collect (\<^bold>g [^] fst yx), s_act2)) rsample_pair_nn) + (return_spmf ((None, CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collect \, s_act2))" +\ \../(Auth1 =a)@(Auth2 =2) \# look\ + | srl_1c_2c: "S_rl (return_spmf (((), CState_Half x, CState_Half y), (auth.State_Store (\<^bold>g [^] x), s_act1), auth.State_Collect (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collect \, s_act2))" + | srl_2c_2c: "S_rl (return_spmf (((), CState_Half x, CState_Half y), (auth.State_Collect (\<^bold>g [^] x), s_act1), auth.State_Collect (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collect \, s_act2))" + | srl_3c_2c: "S_rl (return_spmf (((), CState_Half x, CState_Full (y, z)), (auth.State_Collected, s_act1), auth.State_Collect (\<^bold>g [^] y), s_act2)) + (return_spmf ((Some (z, \<^bold>g [^] x, \<^bold>g [^] y), CState_Half 0, CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collect \, s_act2))" + if "z = (\<^bold>g [^] x) [^] y" +\ \(Auth1 =a)@(Auth2 =3)\ + | srl_1c_3c: "S_rl (return_spmf (((), CState_Full (x, z), CState_Half y), (auth.State_Store (\<^bold>g [^] x), s_act1), auth.State_Collected, s_act2)) + (return_spmf ((Some (z, \<^bold>g [^] x, \<^bold>g [^] y), CState_Full (0, \), CState_Half 0), (auth.State_Store \, s_act1), auth.State_Collected, s_act2))" + if "z = (\<^bold>g [^] y) [^] x" + | srl_2c_3c: "S_rl (return_spmf (((), CState_Full (x, z), CState_Half y), (auth.State_Collect (\<^bold>g [^] x), s_act1), auth.State_Collected, s_act2)) + (return_spmf ((Some (z, \<^bold>g [^] x, \<^bold>g [^] y), CState_Full (0, \), CState_Half 0), (auth.State_Collect \, s_act1), auth.State_Collected, s_act2))" + if "z = (\<^bold>g [^] y) [^] x" + | srl_3c_3c: "S_rl (return_spmf (((), CState_Full (x, z), CState_Full (y, z)), (auth.State_Collected, s_act1), auth.State_Collected, s_act2)) + (return_spmf ((Some (z, \<^bold>g [^] x, \<^bold>g [^] y), CState_Full (0, \), CState_Full (0, \)), (auth.State_Collected, s_act1), auth.State_Collected, s_act2))" + if "z = (\<^bold>g [^] y) [^] x" +\ \(Auth1 =0)@(Auth2 =1')\ + | srl_0_1': "S_rl (return_spmf (((), CState_Void, CState_Half y), (auth.State_Void, s_act1), auth.State_Store (\<^bold>g [^] y), s_act2)) + (map_spmf (\x. ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Store \, s_act2)) rsample_nat)" +\ \(Auth1 =0)@(Auth2 =2')\ + | srl_0_2': "S_rl (return_spmf (((), CState_Void, CState_Half y), (auth.State_Void, s_act1), auth.State_Collect (\<^bold>g [^] y), s_act2)) + (map_spmf (\x. ((Some ((\<^bold>g [^] x) [^] y, \<^bold>g [^] x, \<^bold>g [^] y), CState_Void, CState_Half 0), (auth.State_Void, s_act1), auth.State_Collect \, s_act2)) rsample_nat)" + + +private lemma trac_eq_core_rl: "trace_core_eq real_core' (basic_core DH0_sample) + (UNIV <+> UNIV) ((UNIV <+> UNIV <+> UNIV) <+> UNIV <+> UNIV <+> UNIV) ((UNIV <+> UNIV) <+> UNIV <+> UNIV) + (return_spmf real_s_core') (return_spmf basic_core_sinit)" +proof - + have power_commute: "(\<^bold>g [^] x) [^] (y :: nat) = (\<^bold>g [^] y) [^] (x :: nat)" for x y + by (simp add: nat_pow_pow mult.commute) + + have [simp]: "weight_spmf rsample_nat = 1" + by (simp add: finite_carrier order_gt_0_iff_finite) + + have [simp]: "mk_lossless rsample_nat = rsample_nat" + by (simp add: mk_lossless_def) + + have [simp]: "mk_lossless rsample_pair_nn = rsample_pair_nn" + by (simp add: mk_lossless_def) + + note [simp] = basic_core_oracle_usr_def basic_core_helper_def + exec_gpv_bind spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def Let_def split_def + + show ?thesis + apply (rule trace_core_eq_simI_upto[where S=S_rl]) + subgoal Init_OK + by (simp add: real_s_core'_def srl_0_0) + subgoal POut_OK for s_l s_r query + apply (cases query) + subgoal for e_auth1 by (cases e_auth1; erule S_rl.cases; auto simp add: map_spmf_conv_bind_spmf[symmetric] split!: if_splits) + subgoal for e_auth2 by (cases e_auth2; erule S_rl.cases; auto simp add: map_spmf_conv_bind_spmf[symmetric] split!: if_splits) + done + subgoal PState_OK for s_l s_r query + apply (cases query) + subgoal for e_auth1 by(cases e_auth1; erule S_rl.cases; auto simp add: map_spmf_conv_bind_spmf[symmetric] split!: if_splits intro: S_rl.intros trace_eq_simcl.base) + subgoal for e_auth2 by (cases e_auth2; erule S_rl.cases; auto simp add: map_spmf_conv_bind_spmf[symmetric] split!: if_splits intro: S_rl.intros trace_eq_simcl.base) + done + subgoal AOut_OK for sl sr q + apply (cases q) + subgoal for q_auth1 + apply (cases q_auth1) + subgoal for q_drop by (erule S_rl.cases; simp) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look by(erule S_rl.cases; auto simp add: DH0_sample_def pair_spmf_alt_def) + subgoal for q_fedit by (cases q_fedit; erule S_rl.cases; auto simp add: DH0_sample_def pair_spmf_alt_def) + done + done + subgoal for q_auth2 + apply (cases q_auth2) + subgoal for q_drop by (erule S_rl.cases; simp) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look by(erule S_rl.cases; auto simp add: DH0_sample_def pair_spmf_alt_def) + subgoal for q_fedit by (cases q_fedit; erule S_rl.cases; auto simp add: DH0_sample_def pair_spmf_alt_def) + done + done + done + subgoal AState_OK for sl sr q s1 s2 s1' s2' oa + apply (cases q) + subgoal for q_auth1 + apply (cases q_auth1) + subgoal for q_drop by (erule S_rl.cases; simp) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + proof (erule S_rl.cases, goal_cases) + case (2 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_0}\ + then show ?case + apply(cases s1') + apply (clarsimp simp add: DH0_sample_def) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by(subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl.base S_rl.intros) + next + case (4 x s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1'_0}\ + then show ?case + by(auto simp add: DH0_sample_def map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S_rl.intros) + next + case (7 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_1}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def pair_spmf_alt_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by (subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl_map S_rl.intros) + next + case (13 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_2}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def pair_spmf_alt_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_pair_spmf1[unfolded map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by(subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl_map S_rl.intros) + qed (auto intro: S_rl.intros) + subgoal for q_fedit + apply (cases q_fedit) + by (erule S_rl.cases, goal_cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base intro: S_rl.intros) + done + done + subgoal for q_auth2 + apply (cases q_auth2) + subgoal for q_drop by (erule S_rl.cases; simp) + subgoal for q_lfe + apply (cases q_lfe) + subgoal for q_look + proof (erule S_rl.cases, goal_cases) + case (6 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_1}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply(subst cond_spmf_fst_pair_spmf1[simplified map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by(subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl.base S_rl.intros) + next + case (7 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_1}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply(subst (1 2) cond_spmf_fst_pair_spmf1[simplified map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by(subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl_map S_rl.intros) + next + case (8 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2_1}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(simp add: map_spmf_conv_bind_spmf[symmetric]) + apply(subst (1 2) cond_spmf_fst_pair_spmf1[simplified map_prod_def split_def]) + apply(simp) + apply(subst (1 2) cond_spmf_fst_map_Pair1; simp add: inj_on_def) + by(subst (1 2 3 4) inv_into_f_f; simp add: inj_on_def trace_eq_simcl_map S_rl.intros) + next + case (21 y s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_1'}\ + then show ?case + by(auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base intro: S_rl.intros) + qed (auto intro: S_rl.intros) + subgoal for q_fedit + apply (cases q_fedit) + by (erule S_rl.cases, goal_cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base intro: S_rl.intros) + done + done + done + subgoal UOut_OK for sl sr q + apply (cases q) + subgoal for q_usr + apply (cases q_usr) + subgoal for q_alice by (erule S_rl.cases; simp add: DH0_sample_def pair_spmf_alt_def power_commute) + subgoal for q_bob by (erule S_rl.cases; auto simp add: bind_bind_conv_pair_spmf apfst_def DH0_sample_def power_commute split!: if_split) + done + subgoal for q_act + apply (cases q_act) + subgoal for q_alice + by (erule S_rl.cases; auto simp add: left_gpv_bind_gpv exec_gpv_parallel_oracle_left map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf intro!: bind_spmf_cong) + subgoal for q_bob + by (erule S_rl.cases; auto simp add: right_gpv_bind_gpv exec_gpv_parallel_oracle_right map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf intro!: bind_spmf_cong) + done + done + subgoal UState_OK for sl sr q + apply (cases q) + subgoal for q_usr + apply (cases q_usr) + subgoal for q_alice + proof (erule S_rl.cases, goal_cases) + case (13 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_2}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def pair_spmf_alt_def) + apply(subst (1) bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(subst (1 2) cond_spmf_fst_bind) + by (auto simp add: power_commute intro!: trace_eq_simcl_bind S_rl.intros) + next + case (14 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2_2}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def pair_spmf_alt_def) + apply(subst (1) bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf) + apply(subst (1 2) cond_spmf_fst_bind) + by (auto simp add: power_commute intro!: trace_eq_simcl_bind S_rl.intros) + qed (auto intro: S_rl.intros) + subgoal for q_bob + proof (erule S_rl.cases, goal_cases) + case (8 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2_1}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf power_commute) + apply(subst (1 2) cond_spmf_fst_bind) + by (auto simp add: power_commute intro!: trace_eq_simcl_bind S_rl.intros) + next + case (14 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2_2}\ + then show ?case + apply(clarsimp simp add: DH0_sample_def) + apply(subst bind_commute_spmf) + apply(simp add: bind_bind_conv_pair_spmf power_commute) + apply(subst (1 2) cond_spmf_fst_bind) + by (auto simp add: power_commute intro!: trace_eq_simcl_bind S_rl.intros) + qed (auto simp add: power_commute intro: S_rl.intros) + done + subgoal for q_act + apply (cases q_act) + subgoal for a_alice + proof (erule S_rl.cases, goal_cases) + case (1 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_0}\ + then show ?case + apply (simp add: left_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (6 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_1}\ + then show ?case + apply (simp add: left_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + apply (subst bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (12 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_2}\ + then show ?case + apply (simp add: left_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + apply (subst bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (21 y s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_2'}\ + then show ?case + apply (simp add: left_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl_map S_rl.intros) + next + case (22 y s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_1'}\ + then show ?case + apply (simp add: left_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl_map S_rl.intros) + qed (simp_all split!: if_splits) + subgoal for a_bob + proof (erule S_rl.cases, goal_cases) + case (1 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_0_0}\ + then show ?case + apply(clarsimp simp add: right_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (2 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1_0}\ + then show ?case + apply(clarsimp simp add: right_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + apply (subst bind_commute_spmf, subst bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (3 s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2_0}\ + then show ?case + apply(clarsimp simp add: right_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + apply (subst bind_commute_spmf, subst bind_bind_conv_pair_spmf) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl.base S_rl.intros) + next + case (4 x s_act1 s_act2) \ \Corresponds to @{thm [source] srl_1'_0}\ + then show ?case + apply(clarsimp simp add: right_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl_map S_rl.intros) + next + case (5 x s_act1 s_act2) \ \Corresponds to @{thm [source] srl_2'_0}\ + then show ?case + apply(clarsimp simp add: right_gpv_bind_gpv pair_spmf_alt_def map_gpv_bind_gpv gpv.map_id map_gpv'_bind_gpv map'_lift_spmf split!: if_splits) + by (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro: trace_eq_simcl_map S_rl.intros) + qed (simp_all split!: if_splits) + done + done + done +qed + +lemma trace_eq_fuse_rl: "UNIV \\<^sub>R 1\<^sub>C |\<^sub>= rassocl\<^sub>C \ RES (fused_resource.fuse real_core' real_rest') (real_s_core', real_s_rest') + \ RES (fused_resource.fuse (lazy_core DH0_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit)" +proof - + have fact1: "UNIV \\<^sub>R 1\<^sub>C |\<^sub>= rassocl\<^sub>C \ RES (fused_resource.fuse (basic_core DH0_sample) basic_rest) (basic_core_sinit, basic_rest_sinit) \ + RES (fused_resource.fuse (lazy_core DH0_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit)" + proof - + have [simp]: "\_full \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ \_full) \res RES (fused_resource.fuse (basic_core DH0_sample) basic_rest) (basic_core_sinit, basic_rest_sinit) \" for s + apply (rule WT_resource_of_oracle, rule WT_calleeI) + by (case_tac call, rename_tac [!] x, case_tac [!] x, rename_tac [!] y, case_tac [!] y) + (auto simp add: fused_resource.fuse.simps parallel_eoracle_def) + + note [simp] = exec_gpv_bind spmf.map_comp o_def map_bind_spmf bind_map_spmf bind_spmf_const + + show ?thesis + apply(subst attach_wiring_resource_of_oracle) + apply(rule wiring_parallel_converter2 wiring_id_converter[where \=\_full] wiring_rassocl[of \_full \_full \_full])+ + apply simp_all + apply (rule eq_resource_on_resource_of_oracleI[where S="(=)"]) + apply(simp_all add: eq_on_def relator_eq) + apply(rule ext)+ + apply(subst fuse_ishift_core_to_rest[where core="basic_core DH0_sample" and rest=basic_rest and core'="lazy_core DH0_sample" and + rest'=lazy_rest and fn=basic_core_helper and h_out="map_sum (\_. Out_Activation_Alice) (\_. Out_Activation_Bob)", simplified]) + apply (simp_all add: lazy_rest_def) + apply(fold apply_comp_wiring) + by (simp add: comp_wiring_def parallel2_wiring_def split_def sum.map_comp lassocr\<^sub>w_def rassocl\<^sub>w_def id_def[symmetric] sum.map_id) + qed + + have fact2: "UNIV \\<^sub>R 1\<^sub>C |\<^sub>= rassocl\<^sub>C \ RES (fused_resource.fuse real_core' real_rest') (real_s_core', real_s_rest') \ + 1\<^sub>C |\<^sub>= rassocl\<^sub>C \ RES (fused_resource.fuse (basic_core DH0_sample) basic_rest) (basic_core_sinit, basic_rest_sinit)" + (is "_ \\<^sub>R _ \ RES ?L ?s_l \ _ \ RES ?R ?s_r") proof - + have [simp]: "trace_rest_eq basic_rest basic_rest UNIV UNIV s s" for s + by (rule rel_rest'_into_trace_rest_eq[where S="(=)" and M="(=)"]) (simp_all add: eq_onp_def rel_rest'_eq) + have [simp]: "\_full \\<^sub>\ \_full \c callee_of_rest basic_rest s \" for s + unfolding callee_of_core_def by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + have [simp]: "\_full \\<^sub>\ (\_full \\<^sub>\ \_full) \c callee_of_core (basic_core DH0_sample) s \" for s + unfolding callee_of_core_def by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + have [simp]: "\_full \\<^sub>\ (\_full \\<^sub>\ \_full) \c callee_of_core real_core' s \" for s + unfolding callee_of_core_def by (rule WT_calleeI) (cases s, case_tac call, rename_tac [!] x, case_tac [!] x, auto) + + have loc[simplified]: "((UNIV <+> UNIV) <+> UNIV <+> UNIV) \\<^sub>C ?L(?s_l) \ ?R(?s_r)" + by (rule fuse_trace_eq[where \E=\_full and \CA=\_full and \CU=\_full and \RA=\_full and \RU=\_full, simplified outs_plus_\ outs_\_full]) + (simp_all add: real_rest'_def real_s_rest'_def trac_eq_core_rl[simplified]) + + show ?thesis + apply (rule attach_trace_eq'[where \=\_full and \'=\_full, simplified outs_plus_\ outs_\_full]) + apply (subst trace_eq'_resource_of_oracle, rule loc[simplified]) + by (simp_all add: WT_converter_\_full) + qed + + show ?thesis using fact2[simplified eq_resource_on_UNIV_D[OF fact1]] by blast +qed + +lemma connect_real: "connect D (obsf_resource real_resource) = connect D (obsf_resource (RES (fused_resource.fuse (lazy_core DH0_sample) lazy_rest) (basic_core_sinit, basic_rest_sinit)))" +proof - + have [simp]: "\_full \res real_resource \" + proof - + have [simp]: "\_full \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ \_full) \res RES (fused_resource.fuse real_core' real_rest') (real_s_core', real_s_rest') \" + apply (rule WT_resource_of_oracle) + apply (rule WT_calleeI) + subgoal for s q + apply (cases s, cases q, rename_tac [!] x, case_tac [!] x) + prefer 3 + subgoal for s_cnv_core _ _ _ _ y + apply (cases s_cnv_core, rename_tac s_cnvs s_auth1 s_kern2 s_shell2) + apply (case_tac s_auth1, rename_tac s_kern1 s_shell1) + apply (case_tac s_cnvs, rename_tac su s_cnv1 s_cnv2) + apply (cases y, rename_tac [!] z, case_tac [!] z, rename_tac [!] query) + apply (auto simp add: fused_resource.fuse.simps split_def apfst_def) + apply(case_tac "(s_cnv1, Inl query)" rule: alice_callee.cases; auto split!: sum.splits auth.ousr_bob.splits simp add: Let_def o_def) + apply(case_tac "(s_cnv2, Inl query)" rule: bob_callee.cases; auto split!: sum.splits auth.ousr_bob.splits simp add: Let_def o_def) + apply(case_tac "(s_cnv1, Inr query)" rule: alice_callee.cases; auto split!: sum.splits + simp add: Let_def o_def map_gpv_bind_gpv left_gpv_bind_gpv map_gpv'_bind_gpv exec_gpv_bind) + apply(case_tac "(s_cnv2, Inr query)" rule: bob_callee.cases; auto split!: sum.splits + simp add: Let_def o_def map_gpv_bind_gpv right_gpv_bind_gpv map_gpv'_bind_gpv exec_gpv_bind) + done + by (auto simp add: fused_resource.fuse.simps) + done + + show ?thesis + unfolding attach_real + apply (rule WT_resource_attach[where \'="\_full \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ \_full)"]) + apply (rule WT_converter_mono[of "\_full \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_full))" "\_full \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ \_full)"]) + apply (rule WT_converter_parallel_converter2) + apply (rule WT_intro)+ + by (simp_all add: \_full_le_plus_\) + qed + + show ?thesis + using trace_eq_obsf_resourceI[OF trace_eq_fuse_rl, folded attach_real] + by (rule connect_cong_trace[where A="UNIV" and \=\_full]) + (auto intro!: WT_obsf_resource[where \=\_full, simplified exception_\_full]) +qed + +end + +end + +end + +subsection \Concrete security\ + +context diffie_hellman begin + +context + fixes + auth1_rest :: "('auth1_s_rest, auth.event, 'auth1_iadv_rest, 'auth1_iusr_rest, 'auth1_oadv_rest, 'auth1_ousr_rest) rest_wstate" and + auth2_rest :: "('auth2_s_rest, auth.event, 'auth2_iadv_rest, 'auth2_iusr_rest, 'auth2_oadv_rest, 'auth2_ousr_rest) rest_wstate" and + \_adv_rest1 and \_adv_rest2 and \_usr_rest1 and \_usr_rest2 and I_auth1_rest and I_auth2_rest + assumes + WT_auth1_rest [WT_intro]: "WT_rest \_adv_rest1 \_usr_rest1 I_auth1_rest auth1_rest" and + WT_auth2_rest [WT_intro]: "WT_rest \_adv_rest2 \_usr_rest2 I_auth2_rest auth2_rest" +begin + +theorem secure: + defines "\_real \ ((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` (carrier \)) UNIV)) \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` (carrier \)) UNIV))) \\<^sub>\ (\_adv_rest1 \\<^sub>\ \_adv_rest2)" + and "\_common \ (\_uniform UNIV (key.Out_Alice ` carrier \) \\<^sub>\ \_uniform UNIV (key.Out_Bob ` carrier \)) \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ (\_usr_rest1 \\<^sub>\ \_usr_rest2))" + and "\_ideal \ \_full \\<^sub>\ (\_full \\<^sub>\ (\_adv_rest1 \\<^sub>\ \_adv_rest2))" + shows "constructive_security_obsf + (real_resource TYPE(_) TYPE(_) auth1_rest auth2_rest) + (key.resource (ideal_rest auth1_rest auth2_rest)) + (let sim = CNV sim_callee None in ((sim |\<^sub>= 1\<^sub>C ) \ lassocr\<^sub>C)) + \_real \_ideal \_common \ + (ddh.advantage \ (DH_adversary TYPE(_) TYPE(_) auth1_rest auth2_rest \))" +proof + let ?sim = "(let sim = CNV sim_callee None in ((sim |\<^sub>= 1\<^sub>C ) \ lassocr\<^sub>C))" + + have *[WT_intro]: "(\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` carrier \) UNIV)) \\<^sub>\ + (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` carrier \) UNIV)), \_full \\<^sub>\ \_full \\<^sub>C CNV sim_callee s \" for s + apply (rule WT_converter_of_callee, simp_all) + apply (rename_tac s q r s', case_tac "(s, q)" rule: sim_callee.cases) + by (auto split: if_splits option.splits) + + show "\_real \\<^sub>\ \_common \res real_resource TYPE(_) TYPE(_) auth1_rest auth2_rest \" + proof - + have [WT_intro]: "\_uniform UNIV (key.Out_Alice ` carrier \) \\<^sub>\ \_full, \_uniform (auth.Inp_Send ` carrier \) UNIV \\<^sub>\ \_uniform UNIV (auth.Out_Recv ` carrier \) \\<^sub>C CNV alice_callee CState_Void \" + apply (rule WT_converter_of_callee_invar[where I="pred_cstate (\x. x \ carrier \)"]) + subgoal for s q by (cases "(s, q)" rule: alice_callee.cases) (auto simp add: Let_def split: auth.ousr_bob.splits) + subgoal for s q by (cases "(s, q)" rule: alice_callee.cases) (auto split: if_split_asm auth.ousr_bob.splits simp add: Let_def) + subgoal by simp + done + + have [WT_intro]: "\_uniform UNIV (key.Out_Bob ` carrier \) \\<^sub>\ \_full, \_uniform UNIV (auth.Out_Recv ` carrier \) \\<^sub>\ \_uniform (auth.Inp_Send ` carrier \) UNIV \\<^sub>C CNV bob_callee CState_Void \" + apply (rule WT_converter_of_callee_invar[where I="pred_cstate (\x. x \ carrier \)"]) + subgoal for s q by (cases "(s, q)" rule: bob_callee.cases) (auto simp add: Let_def split: auth.ousr_bob.splits) + subgoal for s q by (cases "(s, q)" rule: bob_callee.cases) (auto simp add: Let_def split: auth.ousr_bob.splits) + subgoal by simp + done + + show ?thesis + unfolding \_real_def \_common_def real_resource_def Let_def fused_wiring_def + by (rule WT_intro)+ + qed + + show "\_ideal \\<^sub>\ \_common \res key.resource (ideal_rest auth1_rest auth2_rest) \" + unfolding \_ideal_def \_common_def key.resource_def + apply(rule callee_invariant_on.WT_resource_of_oracle[where I="\((kernel, _), _, s12). key.set_s_kernel kernel \ carrier \ \ pred_prod I_auth1_rest I_auth2_rest s12"]; (simp add: WT_restD[OF WT_auth1_rest] WT_restD[OF WT_auth2_rest])?) + apply unfold_locales + subgoal for s q + apply (cases "(ideal_rest auth1_rest auth2_rest, s, q)" rule: key.fuse.cases; clarsimp split: if_split_asm) + apply (auto simp add: translate_eoracle_def parallel_eoracle_def plus_eoracle_def) + apply(auto dest: WT_restD_rfunc_adv[OF WT_auth1_rest] WT_restD_rfunc_adv[OF WT_auth2_rest] + WT_restD_rfunc_usr[OF WT_auth1_rest] WT_restD_rfunc_usr[OF WT_auth2_rest] key.foldl_poke_invar) + apply(auto dest!: key.foldl_poke_invar split: plus_oracle_split_asm) + done + subgoal for s + apply(rule WT_calleeI) + subgoal for x y s' + apply(auto simp add: translate_eoracle_def parallel_eoracle_def plus_eoracle_def) + apply(auto dest: WT_restD_rfunc_adv[OF WT_auth1_rest] WT_restD_rfunc_adv[OF WT_auth2_rest] + WT_restD_rfunc_usr[OF WT_auth1_rest] WT_restD_rfunc_usr[OF WT_auth2_rest] split: if_split_asm) + apply(case_tac xa) + apply auto + done + done + done + + show "\_real, \_ideal \\<^sub>C ?sim \" + unfolding \_real_def \_ideal_def Let_def + by(rule WT_intro)+ + + show "pfinite_converter \_real \_ideal ?sim" + proof - + have [pfinite_intro]:"pfinite_converter ((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` carrier \) UNIV)) \\<^sub>\ + (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` carrier \) UNIV))) (\_full \\<^sub>\ \_full) (CNV sim_callee s)" for s + apply(rule raw_converter_invariant.pfinite_converter_of_callee[where I="\_. True"], simp_all) + subgoal + apply (unfold_locales, simp_all) + subgoal for s1 s2 + apply (case_tac "(s1, s2)" rule: sim_callee.cases) + by (auto simp add: id_def split!: sum.splits if_splits option.splits) + done + subgoal for s2 s1 by (case_tac "(s1, s2)" rule: sim_callee.cases) auto + done + + show ?thesis + unfolding \_real_def \_ideal_def Let_def + by (rule pfinite_intro | rule WT_intro)+ + qed + + show "0 \ ddh.advantage \ (diffie_hellman.DH_adversary \ auth1_rest auth2_rest \)" + by(simp add: ddh.advantage_def) + + assume WT [WT_intro]: "exception_\ (\_real \\<^sub>\ \_common) \g \ \" + show "advantage \ (obsf_resource (?sim |\<^sub>= 1\<^sub>C \ key.resource (ideal_rest auth1_rest auth2_rest))) (obsf_resource (real_resource TYPE(_) TYPE(_) auth1_rest auth2_rest)) \ ddh.advantage \ (diffie_hellman.DH_adversary \ auth1_rest auth2_rest \)" + proof - + have id_split[unfolded Let_def]: "connect \ (obsf_resource (?sim |\<^sub>= 1\<^sub>C \ key.resource (ideal_rest auth1_rest auth2_rest))) = + connect \ (obsf_resource (?sim |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ key.resource (ideal_rest auth1_rest auth2_rest)))" (is "connect _ ?L = connect _ ?R") + proof - + note [unfolded \_ideal_def, WT_intro] = \\_real, \_ideal \\<^sub>C ?sim \\ + note [unfolded \_ideal_def, WT_intro] = \\_ideal \\<^sub>\ \_common \res key.resource (ideal_rest auth1_rest auth2_rest) \\ + + have [WT_intro]: "WT_rest (\_full \\<^sub>\ (\_adv_rest1 \\<^sub>\ \_adv_rest2)) (\_full \\<^sub>\ (\_usr_rest1 \\<^sub>\ \_usr_rest2)) (\(_, s12). pred_prod I_auth1_rest I_auth2_rest s12) + (ideal_rest auth1_rest auth2_rest)" + apply (rule WT_rest.intros; simp) + subgoal for s q + apply (cases s, case_tac q, rename_tac [2] x, case_tac [2] x) + apply (auto simp add: translate_eoracle_def parallel_eoracle_def) + using WT_restD_rfunc_adv[OF WT_auth1_rest] WT_restD_rfunc_adv[OF WT_auth2_rest] by fastforce+ + subgoal for s q + apply (cases s, case_tac q, rename_tac [2] x, case_tac [2] x) + apply (auto simp add: translate_eoracle_def parallel_eoracle_def plus_eoracle_def) + using WT_restD_rfunc_usr[OF WT_auth1_rest] WT_restD_rfunc_usr[OF WT_auth2_rest] by fastforce+ + subgoal by(simp add: WT_restD[OF WT_auth1_rest] WT_restD[OF WT_auth2_rest]) + done + + have *: "outs_\ (exception_\ (\_real \\<^sub>\ \_common)) \\<^sub>R ?L \ ?R" + apply (rule obsf_resource_eq_\_cong) + apply (rule eq_\_attach_on') + apply (rule WT_intro | simp)+ + apply(rule parallel_converter2_eq_\_cong) + apply(rule eq_\_converter_reflI) + apply (rule \\_real, \_ideal \\<^sub>C ?sim \\[unfolded assms Let_def]) + apply (rule eq_\_converter_sym) + apply (rule parallel_converter2_id_id) + by (auto simp add: \_real_def \_common_def) + + show ?thesis + by (rule * connect_eq_resource_cong WT_intro)+ + qed + + show ?thesis + unfolding advantage_def Let_def id_split + unfolding Let_def connect_real connect_ideal[unfolded ideal_resource_def Let_def] reduction[unfolded advantage_def] .. + qed +qed + +end + +end + +subsection \Asymptotic security\ + +locale diffie_hellman' = + fixes \ :: "security \ 'grp cyclic_group" + assumes diffie_hellman [locale_witness]: "\\. diffie_hellman (\ \)" +begin + +sublocale diffie_hellman "\ \" for \ .. + +definition real_resource' where "real_resource' rest1 rest2 \ = real_resource TYPE(_) TYPE(_) \ (rest1 \) (rest2 \)" +definition ideal_resource' where "ideal_resource' rest1 rest2 \ = key.resource \ (ideal_rest (rest1 \) (rest2 \))" +definition sim' where "sim' \ = (let sim = CNV (sim_callee \) None in ((sim |\<^sub>= 1\<^sub>C ) \ lassocr\<^sub>C))" + +context + fixes + auth1_rest :: "nat \ ('auth1_s_rest, auth.event, 'auth1_iadv_rest, 'auth1_iusr_rest, 'auth1_oadv_rest, 'auth1_ousr_rest) rest_wstate" and + auth2_rest :: "nat \ ('auth2_s_rest, auth.event, 'auth2_iadv_rest, 'auth2_iusr_rest, 'auth2_oadv_rest, 'auth2_ousr_rest) rest_wstate" and + \_adv_rest1 and \_adv_rest2 and \_usr_rest1 and \_usr_rest2 and I_auth1_rest and I_auth2_rest + assumes + WT_auth1_rest: "\\. WT_rest (\_adv_rest1 \) (\_usr_rest1 \) (I_auth1_rest \) (auth1_rest \)" and + WT_auth2_rest: "\\. WT_rest (\_adv_rest2 \) (\_usr_rest2 \) (I_auth2_rest \) (auth2_rest \)" +begin + +theorem secure: + defines "\_real \ \\. ((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` (carrier (\ \))) UNIV)) \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (auth.Inp_Fedit ` (carrier (\ \))) UNIV))) \\<^sub>\ (\_adv_rest1 \ \\<^sub>\ \_adv_rest2 \)" + and "\_common \ \\. (\_uniform UNIV (key.Out_Alice ` carrier (\ \)) \\<^sub>\ \_uniform UNIV (key.Out_Bob ` carrier (\ \))) \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ (\_usr_rest1 \ \\<^sub>\ \_usr_rest2 \))" + and "\_ideal \ \\. \_full \\<^sub>\ (\_full \\<^sub>\ (\_adv_rest1 \ \\<^sub>\ \_adv_rest2 \))" +assumes DDH: "negligible (\\. ddh.advantage (\ \) (DH_adversary TYPE(_) TYPE(_) \ (auth1_rest \) (auth2_rest \) (\ \)))" + shows "constructive_security_obsf' (real_resource' auth1_rest auth2_rest) (ideal_resource' auth1_rest auth2_rest) sim' \_real \_ideal \_common \" +proof(rule constructive_security_obsf'I) + show "constructive_security_obsf (real_resource' auth1_rest auth2_rest \) + (ideal_resource' auth1_rest auth2_rest \) (sim' \) (\_real \) (\_ideal \) (\_common \) + (\ \) (ddh.advantage (\ \) (DH_adversary TYPE(_) TYPE(_) \ (auth1_rest \) (auth2_rest \) (\ \)))" for \ + unfolding real_resource'_def ideal_resource'_def sim'_def \_real_def \_common_def \_ideal_def + by(rule secure)(rule WT_auth1_rest WT_auth2_rest)+ +qed(rule DDH) + +end + +end + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Constructions/One_Time_Pad.thy b/thys/Constructive_Cryptography_CM/Constructions/One_Time_Pad.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Constructions/One_Time_Pad.thy @@ -0,0 +1,911 @@ +theory One_Time_Pad + imports + Sigma_Commit_Crypto.Xor + "../Asymptotic_Security" + "../Construction_Utility" + "../Specifications/Key" + "../Specifications/Channel" +begin + +section \One-time-pad construction\ + +locale one_time_pad = + key: ideal_key "carrier \" + + auth: ideal_channel "id :: 'msg \ 'msg" False + + sec: ideal_channel "\_ :: 'msg. carrier \" False + + boolean_algebra \ + for + \ :: "('msg, 'more) boolean_algebra_scheme" (structure) + + assumes + nempty_carrier: "carrier \ \ {}" and + finite_carrier: "finite (carrier \)" +begin + + +subsection \Defining user callees\ + +definition enc_callee :: " unit \ 'msg sec.iusr_alice + \ (sec.ousr_alice \ unit, key.iusr_alice + 'msg sec.iusr_alice, 'msg key.ousr_alice + auth.ousr_alice) gpv" + where + "enc_callee \ stateless_callee (\inp. case inp of sec.Inp_Send msg \ + if msg \ carrier \ then + Pause + (Inl key.Inp_Alice) + (\kout. case projl kout of key.Out_Alice key \ + let cipher = key \ msg in + Pause (Inr (auth.Inp_Send cipher)) (\_. Done sec.Out_Send)) + else + Fail)" + +definition dec_callee :: " unit \ sec.iusr_bob + \ ('msg sec.ousr_bob \ unit, key.iusr_bob + auth.iusr_bob, 'msg key.ousr_bob + 'msg auth.ousr_bob) gpv" + where + "dec_callee \ stateless_callee (\_. + Pause + (Inr auth.Inp_Recv) + (\cout. case cout of + Inr (auth.Out_Recv cipher) \ + Pause + (Inl key.Inp_Bob) + (\kout. case projl kout of key.Out_Bob key \ + Done (sec.Out_Recv (key \ cipher))) + | _ \ Fail))" + + +subsection \Defining adversary converter\ + +type_synonym 'msg' astate = "'msg' option" + +definition look_callee :: "'msg astate \ sec.iadv_look + \ ('msg sec.oadv_look \ 'msg astate, sec.iadv_look, 'msg set sec.oadv_look) gpv" + where + "look_callee \ \state inp. + Pause + sec.Inp_Look + (\cout. case cout of + sec.Out_Look msg_set \ + (case state of + None \ do { + msg \ lift_spmf (spmf_of_set (msg_set)); + Done (auth.Out_Look msg, Some msg) } + | Some msg \ Done (auth.Out_Look msg, Some msg)))" + +definition sim :: " + (key.iadv + auth.iadv_drop + auth.iadv_look + 'msg auth.iadv_fedit, + key.oadv + auth.oadv_drop + 'msg auth.oadv_look + auth.oadv_fedit, + sec.iadv_drop + sec.iadv_look + 'msg sec.iadv_fedit, + sec.oadv_drop + 'msg set sec.oadv_look + sec.oadv_fedit) converter" + where + "sim \ + let look_converter = converter_of_callee look_callee None in + ldummy_converter (\_. key.Out_Adversary) (1\<^sub>C |\<^sub>= look_converter |\<^sub>= 1\<^sub>C)" + + +subsection \Defining event-translator\ + +type_synonym estate = "bool \ (key.party + auth.party) set" + +abbreviation einit :: estate + where + "einit \ (False, {})" + +definition sec_party_of_key_party :: "key.party \ sec.party" + where + "sec_party_of_key_party \ key.case_party sec.Alice sec.Bob" + +abbreviation etran_base_helper :: "estate \ key.party + auth.party \ sec.event list" + where + "etran_base_helper \ (\(s_flg, s_kap) item. + let sp_of = case_sum sec_party_of_key_party id in + let se_of = (\chk out. if s_flg \ chk then [out] else []) in + let chk_alice = Inl key.Alice \ s_kap \ Inr auth.Alice \ s_kap in + let chk_bob = Inl key.Bob \ s_kap \ Inr auth.Bob \ s_kap in + sec.case_party + (se_of chk_alice (sec.Event_Shell sec.Alice)) + (se_of chk_bob (sec.Event_Shell sec.Bob)) + (sp_of item))" + +abbreviation etran_base :: "(estate, key.party + auth.party, sec.event list) oracle'" + where + "etran_base \ (\(s_flg, s_kap) item. + let s_kap' = insert item s_kap in + let event = etran_base_helper (s_flg, s_kap') item in + if item \ s_kap then return_spmf (event, s_flg, s_kap') else return_pmf None)" + +fun etran :: "(estate, key.event + auth.event, sec.event list) oracle'" + where + "etran state (Inl (key.Event_Shell party)) = etran_base state (Inl party)" + | "etran (False, s_kap) (Inl key.Event_Kernel) = + (let check_alice = Inl key.Alice \ s_kap \ Inr auth.Alice \ s_kap in + let check_bob = Inl key.Bob \ s_kap \ Inr auth.Bob \ s_kap in + let e_alice = if check_alice then [sec.Event_Shell sec.Alice] else [] in + let e_bob = if check_bob then [sec.Event_Shell sec.Bob] else [] in + return_spmf (e_alice @ e_bob, True, s_kap))" + | "etran state (Inr (auth.Event_Shell party)) = etran_base state (Inr party)" + | "etran _ _ = return_pmf None" + + +subsubsection \Basic lemmas for automated handling of @{term sec_party_of_key_party}\ + +lemma sec_party_of_key_party_simps [simp]: + "sec_party_of_key_party key.Alice = sec.Alice" + "sec_party_of_key_party key.Bob = sec.Bob" + by(simp_all add: sec_party_of_key_party_def) + +lemma sec_party_of_key_party_eq_simps [simp]: + "sec_party_of_key_party p = sec.Alice \ p = key.Alice" + "sec_party_of_key_party p = sec.Bob \ p = key.Bob" + by(simp_all add: sec_party_of_key_party_def split: key.party.split) + +lemma key_case_party_collapse [simp]: "key.case_party x x p = x" + by(simp split: key.party.split) + +lemma sec_case_party_collapse [simp]: "sec.case_party x x p = x" + by(simp split: sec.party.split) + +lemma Alice_in_sec_party_of_key_party [simp]: + "sec.Alice \ sec_party_of_key_party ` P \ key.Alice \ P" + by(auto simp add: sec_party_of_key_party_def split: key.party.splits) + +lemma Bob_in_sec_party_of_key_party [simp]: + "sec.Bob \ sec_party_of_key_party ` P \ key.Bob \ P" + by(auto simp add: sec_party_of_key_party_def split: key.party.splits) + +lemma case_sec_party_of_key_party [simp]: "sec.case_party a b (sec_party_of_key_party x) = key.case_party a b x" + by(simp add: sec_party_of_key_party_def split: sec.party.split key.party.split) + + +subsection \Defining Ideal and Real constructions\ + +context + fixes + key_rest :: "('key_s_rest, key.event, 'key_iadv_rest, 'key_iusr_rest, 'key_oadv_rest, 'key_ousr_rest) rest_wstate" and + auth_rest :: "('auth_s_rest, auth.event, 'auth_iadv_rest, 'auth_iusr_rest, 'auth_oadv_rest, 'auth_ousr_rest) rest_wstate" +begin + +definition ideal_rest + where + "ideal_rest \ translate_rest einit etran (parallel_rest key_rest auth_rest)" + +definition ideal_resource + where + "ideal_resource \ (sim |\<^sub>= 1\<^sub>C) |\<^sub>= 1\<^sub>C |\<^sub>= 1\<^sub>C \ (sec.resource ideal_rest)" + +definition real_resource + where + "real_resource \ attach_c1f22_c1f22 (CNV enc_callee ()) (CNV dec_callee ()) (key.resource key_rest) (auth.resource auth_rest)" + + +subsection \Wiring and simplifying the Ideal construction\ + +definition ideal_s_core' :: "((_ \ 'msg astate \ _) \ _) \ estate \ 'msg sec.state" + where + "ideal_s_core' \ ((((), None, ()), ()), (False, {}), sec.State_Void, {})" + +definition ideal_s_rest' :: "_ \ 'key_s_rest \ 'auth_s_rest" + where + "ideal_s_rest' \ (((), ()), rinit key_rest, rinit auth_rest)" + +primcorec ideal_core' :: "(((unit \ _ \ unit) \ unit) \ _, _, key.iadv + _, _, _, _) core" + where + "cpoke ideal_core' = (\(s_advusr, s_event, s_core) event. do { + (events, s_event') \ (etran s_event event); + s_core' \ foldl_spmf sec.poke (return_spmf s_core) events; + return_spmf (s_advusr, s_event', s_core') + })" + | "cfunc_adv ideal_core' = (\((s_adv, s_usr), s_core) iadv. + let handle_l = (\_. Done (Inl key.Out_Adversary, s_adv)) in + let handle_r = (\qr. map_gpv (map_prod Inr id) id ((1\<^sub>I \\<^sub>I look_callee \\<^sub>I 1\<^sub>I) s_adv qr)) in + map_spmf + (\((oadv, s_adv'), s_core'). (oadv, (s_adv', s_usr), s_core')) + (exec_gpv \sec.iface_adv (case_sum handle_l handle_r iadv) s_core))" + | "cfunc_usr ideal_core' = \\sec.iface_usr" + +primcorec ideal_rest' :: "((unit \ unit) \ _, _, _, _, _, _, _) rest_scheme" + where + "rinit ideal_rest' = (((), ()), rinit key_rest, rinit auth_rest)" + | "rfunc_adv ideal_rest' = \(parallel_eoracle (rfunc_adv key_rest) (rfunc_adv auth_rest))" + | "rfunc_usr ideal_rest' = \(parallel_eoracle (rfunc_usr key_rest) (rfunc_usr auth_rest))" + + +subsubsection \The ideal attachment lemma\ + +lemma attach_ideal: "ideal_resource = RES (fused_resource.fuse ideal_core' ideal_rest') (ideal_s_core', ideal_s_rest')" +proof - + + have fact1: "ideal_rest' = attach_rest 1\<^sub>I 1\<^sub>I (Pair ((), ())) (parallel_rest key_rest auth_rest)" (is "?L = ?R") + proof - + + have "rinit ?L = rinit ?R" + by simp + + moreover have "rfunc_adv ?L = rfunc_adv ?R" + unfolding attach_rest_id_oracle_adv parallel_eoracle_def + by (simp add: extend_state_oracle_def) + + moreover have "rfunc_usr ?L = rfunc_usr ?R" + unfolding attach_rest_id_oracle_usr parallel_eoracle_def + by (simp add: extend_state_oracle_def) + + ultimately show ?thesis + by (coinduction) blast + qed + + have fact2: "ideal_core' = + (let handle_l = (\s ql. Generative_Probabilistic_Value.Done (Inl key.Out_Adversary, s)) in + let handle_r = (\s qr. map_gpv (map_prod Inr id) id ((1\<^sub>I \\<^sub>I look_callee \\<^sub>I 1\<^sub>I) s qr)) in + let tcore = translate_core etran sec.core in + attach_core (\s. case_sum (handle_l s) (handle_r s)) 1\<^sub>I tcore)" (is "?L = ?R") + proof - + + have "cpoke ?L = cpoke ?R" + by (simp add: split_def map_spmf_conv_bind_spmf) + + moreover have "cfunc_adv ?L = cfunc_adv ?R" + unfolding attach_core_def + by (simp add: split_def) + + moreover have "cfunc_usr ?L = cfunc_usr ?R" + unfolding Let_def attach_core_id_oracle_usr + by (clarsimp simp add: extend_state_oracle_def[symmetric]) + + ultimately show ?thesis + by (coinduction) blast + qed + + show ?thesis + unfolding ideal_resource_def sec.resource_def sim_def ideal_rest_def ideal_s_core'_def ideal_s_rest'_def + apply(simp add: conv_callee_parallel_id_right[symmetric, where s'="()"]) + apply(simp add: conv_callee_parallel_id_left[symmetric, where s="()"]) + apply(simp add: ldummy_converter_of_callee) + apply(subst fused_resource_move_translate[of _ einit etran]) + apply(simp add: resource_of_oracle_state_iso) + apply(simp add: iso_swapar_def split_beta ideal_rest_def) + apply(subst (1 2 3) converter_of_callee_id_oracle[symmetric, of "()"]) + apply(subst attach_parallel_fuse'[where f_init="Pair ((), ())"]) + apply(simp add: fact1[symmetric] fact2[symmetric, simplified Let_def]) + done +qed + +subsection \Wiring and simplifying the Real construction\ + +definition real_s_core' :: "_ \ 'msg key.state \ 'msg auth.state" + where + "real_s_core' \ (((), (), ()), (key.PState_Store, {}), (auth.State_Void, {}))" + +definition real_s_rest' + where + "real_s_rest' \ ideal_s_rest'" + +primcorec real_core' :: "((unit \ _) \ _, _, _, _, _, _) core" + where + "cpoke real_core' = (\(s_advusr, s_core) event. + map_spmf (Pair s_advusr) (parallel_handler key.poke auth.poke s_core event))" + | "cfunc_adv real_core' = \(key.iface_adv \\<^sub>O auth.iface_adv)" + | "cfunc_usr real_core' = (\((s_adv, s_usr), s_core) iusr. + let handle_req = lsumr \ map_sum id (rsuml \ map_sum swap_sum id \ lsumr) \ rsuml in + let handle_ret = lsumr \ (map_sum id (rsuml \ (map_sum swap_sum id \ lsumr)) \ rsuml) in + map_spmf + (\((ousr, s_usr'), s_core'). (ousr, (s_adv, s_usr'), s_core')) + (exec_gpv + (key.iface_usr \\<^sub>O auth.iface_usr) + (map_gpv' id handle_req handle_ret ((enc_callee \\<^sub>I dec_callee) s_usr iusr)) s_core))" + +definition real_rest' + where + "real_rest' \ ideal_rest'" + + +subsubsection \The real attachment lemma\ + +private lemma WT_callee_real1: "((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full)) \\<^sub>\ ((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full)) \c + (key.fuse key_rest \\<^sub>O auth.fuse auth_rest) s \" + apply(rule WT_calleeI) + apply(cases s) + apply(case_tac call) + apply(rename_tac [!] x) + apply(case_tac [!] x) + apply(rename_tac [!] y) + apply(case_tac [!] y) + by(auto simp add: fused_resource.fuse.simps) + +private lemma WT_callee_real2: "(\_full \\<^sub>\ \_full) \\<^sub>\ (((\_full \\<^sub>\ \_full) \\<^sub>\ (\_full \\<^sub>\ \_full)) \\<^sub>\ \_full) \c + fused_resource.fuse (parallel_core key.core auth.core) (parallel_rest key_rest auth_rest) s \" + apply(rule WT_calleeI) + apply(cases s) + apply(case_tac call) + apply(rename_tac [!] x) + apply(case_tac [!] x) + apply(rename_tac [!] y) + apply(case_tac [!] y) + apply(rename_tac [5] z) + apply(rename_tac [6] z) + apply(case_tac [5] z) + apply(case_tac [7] z) + by(auto simp add: fused_resource.fuse.simps) + + +lemma attach_real: "real_resource = RES (fused_resource.fuse real_core' real_rest') (real_s_core', real_s_rest')" +proof - + + have fact1: "real_core' = attach_core 1\<^sub>I (attach_wiring_right parallel_wiring\<^sub>w (enc_callee \\<^sub>I dec_callee)) + (parallel_core key.core auth.core) " (is "?L = ?R") + proof- + + have "cpoke ?L = cpoke ?R" + by simp + + moreover have "cfunc_adv ?L = cfunc_adv ?R" + unfolding attach_core_id_oracle_adv + by (simp add: extend_state_oracle_def) + + moreover have "cfunc_usr ?L = cfunc_usr ?R" + unfolding parallel_wiring\<^sub>w_def swap_lassocr\<^sub>w_def swap\<^sub>w_def lassocr\<^sub>w_def rassocl\<^sub>w_def + by (simp add: attach_wiring_right_simps parallel2_wiring_simps comp_wiring_simps) + + ultimately show ?thesis + by (coinduction) blast + qed + + have fact2: "real_rest' = attach_rest 1\<^sub>I 1\<^sub>I (Pair ((), ())) (parallel_rest key_rest auth_rest) " (is "?L = ?R") + proof - + have "rinit ?L = rinit ?R" + unfolding real_rest'_def ideal_rest'_def + by simp + + moreover have "rfunc_adv ?L = rfunc_adv ?R" + unfolding real_rest'_def ideal_rest'_def attach_rest_id_oracle_adv + by (simp add: extend_state_oracle_def) + + moreover have "rfunc_usr ?L = rfunc_usr ?R" + unfolding real_rest'_def ideal_rest'_def attach_rest_id_oracle_usr + by (simp add: extend_state_oracle_def) + + ultimately show ?thesis + by (coinduction) blast + qed + + show ?thesis + unfolding real_resource_def attach_c1f22_c1f22_def wiring_c1r22_c1r22_def key.resource_def auth.resource_def + apply(subst resource_of_parallel_oracle[symmetric]) + apply(subst attach_compose) + apply(subst attach_wiring_resource_of_oracle) + apply(rule wiring_intro) + apply (rule WT_resource_of_oracle[OF WT_callee_real1]) + apply simp + subgoal + apply(subst parallel_oracle_fuse) + apply(subst resource_of_oracle_state_iso) + apply simp + apply(simp add: parallel_state_iso_def) + apply(subst conv_callee_parallel[symmetric]) + apply(subst eq_resource_on_UNIV_iff[symmetric]) + apply(rule eq_resource_on_trans) + apply(rule eq_\_attach_on') + apply (rule WT_resource_of_oracle[OF WT_callee_real2]) + apply(rule parallel_converter2_eq_\_cong) + apply(rule eq_\_converter_reflI) + apply(rule WT_intro)+ + apply(rule parallel_converter2_eq_\_cong) + apply(rule comp_converter_of_callee_wiring) + apply(rule wiring_intro) + apply(subst conv_callee_parallel) + apply(rule WT_intro) + apply (rule WT_converter_of_callee[where \=\_full and \'="\_full \\<^sub>\ \_full"]) + apply (rule WT_gpv_\_mono) + apply (rule WT_gpv_full) + apply (rule \_full_le_plus_\) + apply(rule order_refl) + apply(rule order_refl) + apply (clarsimp simp add: enc_callee_def stateless_callee_def split!: sec.iusr_alice.splits key.ousr_alice.splits) + apply (rule WT_converter_of_callee[where \=\_full and \'="\_full \\<^sub>\ \_full"]) + apply (rule WT_gpv_\_mono) + apply (rule WT_gpv_full) + apply (rule \_full_le_plus_\) + apply(rule order_refl) + apply(rule order_refl) + apply (clarsimp simp add: enc_callee_def stateless_callee_def split!: sec.iusr_alice.splits key.ousr_alice.splits) + apply(subst id_converter_eq_self) + apply(rule order_refl) + apply simp + apply simp + apply(subst eq_resource_on_UNIV_iff) + apply(subst (1 2 3) converter_of_callee_id_oracle[symmetric, of "()"]) + apply(subst attach_parallel_fuse') + apply(simp add: fact1 fact2 real_s_core'_def real_s_rest'_def ideal_s_rest'_def) + done + done +qed + + +subsection \Proving the trace-equivalence of simplified Ideal and Real constructions\ + +context +begin + + +subsubsection \Proving the trace-equivalence of cores\ + +private abbreviation + "a_I \ \(x, y). ((((), x, ()), ()), y)" + +private abbreviation + "a_R \ \x. (((), (), ()), x)" + +abbreviation + "asm_act \ (\flg pset_sec pset_key pset_auth pset_union. + pset_union = pset_key <+> pset_auth \ + (flg \ pset_sec = sec_party_of_key_party ` pset_key \ pset_auth))" + +private inductive S :: "(((_ \ 'msg option \ _) \ _) \ estate \ 'msg sec.state) spmf + \ (_ \ 'msg key.state \ 'msg auth.state) spmf \ bool" + where +\ \(Auth =a)@(Key =0)\ + s_0_0: "S (return_spmf (a_I (None, (False, s_act_ka), sec.State_Void, s_act_s))) + (return_spmf (a_R ((key.PState_Store, s_act_k), auth.State_Void, s_act_a)))" + if "asm_act False s_act_s s_act_k s_act_a s_act_ka" and "s_act_s = {}" +\ \(Auth =a)@(Key =1)\ + | s_0_1: "S (return_spmf (a_I (None, (True, s_act_ka), sec.State_Void, s_act))) + (map_spmf (\key. a_R ((key.State_Store key, s_act_k), auth.State_Void, s_act_a)) (spmf_of_set (carrier \)))" + if "asm_act True s_act s_act_k s_act_a s_act_ka" +\ \../(Auth =a)@(Key =1) \# wl\ + | s_1_1: "S (return_spmf (a_I (None, (True ,s_act_ka), sec.State_Store msg, s_act_s))) + (map_spmf (\key. a_R ((key.State_Store key, s_act_k), auth.State_Store (key \ msg), s_act_a)) (spmf_of_set (carrier \)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "key.Alice \ s_act_k" and "auth.Alice \ s_act_a" and "msg \ carrier \" + | s_2_1: "S (return_spmf (a_I (None, (True ,s_act_ka), sec.State_Collect msg, s_act_s))) + (map_spmf (\key. a_R ((key.State_Store key, s_act_k), auth.State_Collect (key \ msg), s_act_a)) (spmf_of_set (carrier \)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "key.Alice \ s_act_k" and "auth.Alice \ s_act_a" and "msg \ carrier \" + | s_3_1: "S (return_spmf (a_I (None, (True ,s_act_ka), sec.State_Collected, s_act_s))) + (map_spmf (\key. a_R ((key.State_Store key, s_act_k), auth.State_Collected, s_act_a)) (spmf_of_set (carrier \)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "s_act_k = {key.Alice, key.Bob}" and "s_act_a = {auth.Alice, auth.Bob}" +\ \../(Auth =a)@(Key =1) \# look\ + | s_1'_1: "S (return_spmf (a_I (Some (key \ msg), (True ,s_act_ka), sec.State_Store msg, s_act_s))) + (return_spmf (a_R ((key.State_Store key, s_act_k), auth.State_Store (key \ msg), s_act_a)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "key.Alice \ s_act_k" and "auth.Alice \ s_act_a" and "msg \ carrier \" and "key \ carrier \" + | s_2'_1: "S (return_spmf (a_I (Some (key \ msg), (True ,s_act_ka), sec.State_Collect msg, s_act_s))) + (return_spmf (a_R ((key.State_Store key, s_act_k), auth.State_Collect (key \ msg), s_act_a)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "key.Alice \ s_act_k" and "auth.Alice \ s_act_a" and "msg \ carrier \" and "key \ carrier \" + | s_3'_1: "S (return_spmf (a_I (Some (key \ msg), (True ,s_act_ka), sec.State_Collected, s_act_s))) + (return_spmf (a_R ((key.State_Store key, s_act_k), auth.State_Collected, s_act_a)))" + if "asm_act True s_act_s s_act_k s_act_a s_act_ka" and "s_act_k = {key.Alice, key.Bob}" and "s_act_a = {auth.Alice, auth.Bob}" and "msg \ carrier \" and "key \ carrier \" + +private lemma trace_eq_core: "trace_core_eq ideal_core' real_core' + UNIV (UNIV <+> UNIV <+> UNIV <+> (auth.Inp_Fedit ` carrier \)) ((sec.Inp_Send ` carrier \) <+> UNIV) + (return_spmf ideal_s_core') (return_spmf real_s_core')" +proof - + + have inj_xor: "\msg \ carrier \ ; x \ carrier \; y \ carrier \; x \ msg = y \ msg\ \ x = y" for msg x y + by (metis (no_types, hide_lams) local.xor_ac(2) local.xor_left_inverse) + + note [simp] = enc_callee_def dec_callee_def look_callee_def nempty_carrier finite_carrier + exec_gpv_bind spmf.map_comp map_bind_spmf bind_map_spmf bind_spmf_const o_def Let_def + + show ?thesis + apply (rule trace_core_eq_simI_upto[where S=S]) + subgoal Init_OK + by (simp add: ideal_s_core'_def real_s_core'_def S.simps) + subgoal POut_OK for s_i s_r query + apply (cases query) + subgoal for e_key + apply (cases e_key) + subgoal for e_shell by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] split: key.party.splits) + subgoal e_kernel by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric]) + done + subgoal for e_auth + apply (cases e_auth) + subgoal for e_shell + by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] split:auth.party.splits) + done + done + subgoal PState_OK for s_i s_r query + apply (cases query) + subgoal for e_key + apply (cases e_key) + subgoal for e_shell by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S.intros[simplified] split: key.party.splits) + subgoal e_kernel by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] sec_party_of_key_party_def intro!: trace_eq_simcl.base S.intros[simplified] split: key.party.splits) + done + subgoal for e_auth + apply (cases e_auth) + subgoal for e_shell by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S.intros[simplified] split:auth.party.splits) + done + done + subgoal AOut_OK for s_i s_r query + apply (cases query) + subgoal for q_key by (erule S.cases) simp_all + subgoal for q_auth + apply (cases q_auth) + subgoal for q_auth_drop by (erule S.cases) (simp_all add: id_oracle_def) + subgoal for q_auth_lfe + apply (cases q_auth_lfe) + subgoal for q_auth_look + proof (erule S.cases, goal_cases) + case (3 s_act_s s_act_k s_act_a s_act_ka msg) \ \Corresponds to @{thm [source] s_1_1}\ + then show ?case + apply(simp add: exec_gpv_extend_state_oracle exec_gpv_map_gpv_id exec_gpv_plus_oracle_right exec_gpv_plus_oracle_left) + apply (subst one_time_pad[symmetric, of "msg"]) + apply (simp_all add: xor_comm) + apply (rule bind_spmf_cong[OF HOL.refl]) + by (simp add: xor_comm) + qed simp_all + subgoal for q_auth_fedit by (erule S.cases) (auto simp add: id_oracle_def split:auth.iadv_fedit.split) + done + done + done + subgoal AState_OK for s_i s_r query + apply (cases query) + subgoal for q_key by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] intro!: trace_eq_simcl.base S.intros[simplified]) + subgoal for q_auth + apply (cases q_auth) + subgoal for q_auth_drop by (erule S.cases) (auto simp add: id_oracle_def) + subgoal for q_auth_lfe + apply (cases q_auth_lfe) + subgoal for q_auth_look + proof (erule S.cases, goal_cases) + case (3 s_act_s s_act_k s_act_a s_act_ka msg) \ \Corresponds to @{thm [source] s_1_1}\ + then show ?case + apply(simp add: exec_gpv_extend_state_oracle exec_gpv_map_gpv_id exec_gpv_plus_oracle_right exec_gpv_plus_oracle_left) + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric]) + apply (subst (1 2) cond_spmf_fst_map_Pair1; clarsimp simp add: set_spmf_of_set inj_on_def intro: inj_xor) + apply (rule inj_xor, simp_all) + apply(subst (1 2 3) inv_into_f_f) + by (auto simp add: S.simps inj_on_def intro: inj_xor) + qed (auto intro!: trace_eq_simcl.base S.intros[simplified]) + subgoal for q_auth_fedit by (erule S.cases) (auto simp add: map_spmf_conv_bind_spmf[symmetric] id_oracle_def intro!: trace_eq_simcl.base S.intros[simplified]) + done + done + done + subgoal UOut_OK for s_i s_r query + apply (cases query) + subgoal for q_alice + proof (erule S.cases, goal_cases) + case (2 s_act_s s_act_k s_act_a s_act_ka) \ \Corresponds to @{thm [source] s_0_1}\ + then show ?case + apply (cases "auth.Alice \ s_act_a"; cases "key.Alice \ s_act_k") + apply (simp_all add: stateless_callee_def split_def split!: auth.iusr_alice.split) + done + qed (simp_all add: stateless_callee_def split: auth.iusr_alice.split) + subgoal for q_bob + proof (erule S.cases, goal_cases) + case (4 s_act_s s_act_k s_act_a s_act_ka msg) \ \Corresponds to @{thm [source] s_2_1}\ + then show ?case + apply (cases "sec.Bob \ s_act_s") + subgoal + apply (clarsimp simp add: stateless_callee_def) + apply (simp add: spmf_rel_eq[symmetric]) + apply (rule rel_spmf_bindI2) + by simp_all + subgoal by (cases "sec.Bob \ s_act_a") (clarsimp simp add: stateless_callee_def)+ + done + qed (simp_all add: stateless_callee_def) + done + subgoal UState_OK for s_i s_r query + apply (cases query) + subgoal for q_alice + proof (erule S.cases, goal_cases) + case (2 s_act s_act_k s_act_a s_act_ka) \ \Corresponds to @{thm [source] s_0_1}\ + then show ?case + apply (cases "auth.Alice \ s_act_a"; cases "key.Alice \ s_act_k") + subgoal + apply (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] stateless_callee_def split_def split!: auth.iusr_alice.split if_splits) + apply(rule trace_eq_simcl.base) + apply (rule S.intros(3)[simplified]) + by simp_all + by (clarsimp simp add: map_spmf_conv_bind_spmf[symmetric] stateless_callee_def split_def split: auth.iusr_alice.split)+ + qed (auto simp add: stateless_callee_def split: auth.iusr_alice.split_asm) + subgoal for q_bob + proof (erule S.cases, goal_cases) + case (4 s_act_s s_act_k s_act_a s_act_ka msg) \ \Corresponds to @{thm [source] s_2_1}\ + then show ?case + apply (cases "sec.Bob \ s_act_s") + subgoal + apply (clarsimp simp add: stateless_callee_def map_spmf_conv_bind_spmf[symmetric]) + apply (subst map_spmf_of_set_inj_on) + apply (simp_all add: inj_on_def) + apply (subst map_spmf_of_set_inj_on[symmetric]) + apply (simp add: inj_on_def) + apply clarsimp + apply(rule trace_eq_simcl.base) + apply (rule S.intros(5)[simplified]) + apply (simp_all split: sec.party.splits ) + by auto + subgoal by (clarsimp simp add: stateless_callee_def split: if_splits) + done + next + case (7 s_act_s s_act_k s_act_a s_act_ka msg key) \ \Corresponds to @{thm [source] s_2'_1}\ + then show ?case + apply (cases "sec.Bob \ s_act_s") + subgoal + apply (clarsimp simp add: stateless_callee_def map_spmf_conv_bind_spmf[symmetric]) + apply (rule S.intros(8)[simplified]) + apply simp_all + by auto + subgoal by (clarsimp simp add: stateless_callee_def split: if_splits) + done + qed (auto simp add: stateless_callee_def split: auth.iusr_alice.split_asm) + done + done +qed + + +subsubsection \Proving the trace equivalence of fused cores and rests\ + +private definition \_adv_core :: "(key.iadv + 'msg auth.iadv, key.oadv + 'msg auth.oadv) \" + where "\_adv_core \ \_full \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (sec.Inp_Fedit ` (carrier \)) UNIV))" + +private definition \_usr_core :: "('msg sec.iusr, 'msg sec.ousr) \" + where "\_usr_core \ \_uniform (sec.Inp_Send ` (carrier \)) UNIV \\<^sub>\ \_uniform UNIV (sec.Out_Recv ` carrier \)" + +private definition invar_ideal' :: "((_ \ 'msg astate \ _) \ _) \ estate \ 'msg sec.state \ bool" + where "invar_ideal' = pred_prod (pred_prod (pred_prod (\_. True) (pred_prod (pred_option (\x. x \ carrier \)) (\_. True))) (\_. True)) (pred_prod (\_. True) (pred_prod (sec.pred_s_kernel (\x. x \ carrier \)) (\_. True)))" + +private definition invar_real' :: "_ \ ('msg key.s_kernel \ _) \ 'msg sec.s_kernel \ _ \ bool" + where "invar_real' = pred_prod (\_. True) (pred_prod (pred_prod (key.pred_s_kernel (\x. x \ carrier \)) (\_. True)) (pred_prod (sec.pred_s_kernel (\x. x \ carrier \)) (\_. True)))" + +lemma invar_ideal_s_core' [simp]: "invar_ideal' ideal_s_core'" + by(simp add: invar_ideal'_def ideal_s_core'_def) + +lemma invar_real_s_core' [simp]: "invar_real' real_s_core'" + by(simp add: invar_real'_def real_s_core'_def) + +lemma WT_ideal_core' [WT_intro]: "WT_core \_adv_core \_usr_core invar_ideal' ideal_core'" + apply(rule WT_core.intros) + apply + (auto split!: sum.splits option.splits if_split_asm simp add: \_adv_core_def \_usr_core_def exec_gpv_map_gpv_id exec_gpv_extend_state_oracle exec_gpv_plus_oracle_left exec_gpv_plus_oracle_right invar_ideal'_def sec.in_set_spmf_iface_drop sec.in_set_spmf_iface_look sec.in_set_spmf_iface_fedit sec.in_set_spmf_iface_alice sec.in_set_spmf_iface_bob id_oracle_def look_callee_def exec_gpv_bind set_spmf_of_set sec.poke_alt_def foldl_spmf_pair_right) + done + +lemma WT_ideal_rest' [WT_intro]: + assumes "WT_rest \_adv_restk \_usr_restk I_key_rest key_rest" + and "WT_rest \_adv_resta \_usr_resta I_auth_rest auth_rest" + shows "WT_rest (\_adv_restk \\<^sub>\ \_adv_resta) (\_usr_restk \\<^sub>\ \_usr_resta) (\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest) ideal_rest'" + by(rule WT_rest.intros)(fastforce simp add: fused_resource.fuse.simps parallel_eoracle_def dest: WT_restD_rfunc_adv[OF assms(1)] WT_restD_rfunc_adv[OF assms(2)] WT_restD_rfunc_usr[OF assms(1)] WT_restD_rfunc_usr[OF assms(2)] simp add: assms[THEN WT_restD_rinit])+ + + +lemma WT_real_core' [WT_intro]: "WT_core \_adv_core \_usr_core invar_real' real_core'" + apply(rule WT_core.intros) + apply(auto simp add: \_adv_core_def \_usr_core_def enc_callee_def dec_callee_def + stateless_callee_def Let_def exec_gpv_extend_state_oracle exec_gpv_map_gpv' exec_gpv_plus_oracle_left exec_gpv_plus_oracle_right + invar_real'_def in_set_spmf_parallel_handler key.in_set_spmf_poke sec.poke_alt_def auth.in_set_spmf_iface_look auth.in_set_spmf_iface_fedit + sec.in_set_spmf_iface_alice sec.in_set_spmf_iface_bob + split!: key.ousr_alice.splits key.ousr_bob.splits auth.ousr_alice.splits auth.ousr_bob.splits sum.splits if_split_asm) + done + +private lemma trace_eq_sec: + fixes \_adv_restk \_adv_resta \_usr_restk \_usr_resta + defines "outs_adv \ (UNIV <+> UNIV <+> UNIV <+> sec.Inp_Fedit ` carrier \) <+> outs_\ (\_adv_restk \\<^sub>\ \_adv_resta)" + and "outs_usr \ (sec.Inp_Send ` carrier \ <+> UNIV) <+> outs_\ (\_usr_restk \\<^sub>\ \_usr_resta)" + assumes WT_key [WT_intro]: "WT_rest \_adv_restk \_usr_restk I_key_rest key_rest" + and WT_auth [WT_intro]: "WT_rest \_adv_resta \_usr_resta I_auth_rest auth_rest" + shows "(outs_adv <+> outs_usr) \\<^sub>C fused_resource.fuse ideal_core' ideal_rest' ((ideal_s_core', ideal_s_rest')) \ + fused_resource.fuse real_core' real_rest' ((real_s_core', real_s_rest'))" +proof - + define e\_adv_rest :: "(_, _ \ (key.event + auth.event) list) \" + where "e\_adv_rest \ map_\ id (case_sum (map_prod Inl (map Inl)) (map_prod Inr (map Inr))) (e\ \_adv_restk \\<^sub>\ e\ \_adv_resta)" + define e\_usr_rest :: "(_, _ \ (key.event + auth.event) list) \" + where "e\_usr_rest \ map_\ id (case_sum (map_prod Inl (map Inl)) (map_prod Inr (map Inr))) (e\ \_usr_restk \\<^sub>\ e\ \_usr_resta)" + + note I_defs = \_adv_core_def \_usr_core_def + note eI_defs = e\_adv_rest_def e\_usr_rest_def + + have fact1[unfolded outs_plus_\]: + "trace_rest_eq ideal_rest' ideal_rest' (outs_\ (\_adv_restk \\<^sub>\ \_adv_resta)) (outs_\ (\_usr_restk \\<^sub>\ \_usr_resta)) s s" for s + apply(rule rel_rest'_into_trace_rest_eq[where S="(=)" and M="(=)", unfolded eq_onp_def], simp_all) + apply(fold relator_eq) + apply(rule rel_rest'_mono[THEN predicate2D, rotated -1, OF HOL.refl[of ideal_rest', folded relator_eq]]) + by auto + + have fact2 [unfolded eI_defs]: "callee_invariant_on (callee_of_rest ideal_rest') (\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest) (e\_adv_rest \\<^sub>\ e\_usr_rest)" + apply unfold_locales + subgoal for s x y s' + apply(cases "(snd s, x)" rule: parallel_oracle.cases) + apply(auto 4 3 simp add: parallel_eoracle_def eI_defs split!: sum.splits dest: WT_restD(1,2)[OF WT_key] WT_restD(1,2)[OF WT_auth]) + done + subgoal for s + apply(fastforce intro!: WT_calleeI simp add: parallel_eoracle_def eI_defs image_image dest: WT_restD(1,2)[OF WT_key] WT_restD(1,2)[OF WT_auth] intro: rev_image_eqI) + done + done + + have fact3[unfolded I_defs]: "callee_invariant_on (callee_of_core ideal_core') invar_ideal' (\_full \\<^sub>\ (\_adv_core \\<^sub>\ \_usr_core))" + by(rule WT_intro)+ + + have fact4[unfolded I_defs]: "callee_invariant_on (callee_of_core real_core') invar_real' (\_full \\<^sub>\ (\_adv_core \\<^sub>\ \_usr_core))" + by(rule WT_intro)+ + + note nempty_carrier[simp] + show ?thesis using WT_key[THEN WT_restD_rinit] WT_auth[THEN WT_restD_rinit] + apply (simp add: real_rest'_def real_s_rest'_def assms(1, 2)) + thm fuse_trace_eq[where \E=\_full and \CA=\_adv_core and \CU=\_usr_core and \RA=e\_adv_rest and \RU=e\_usr_rest,unfolded eI_defs \_adv_core_def \_usr_core_def,simplified] + apply (rule fuse_trace_eq[where \E=\_full and \CA=\_adv_core and \CU=\_usr_core and \RA=e\_adv_rest and \RU=e\_usr_rest + and ?IR1.0 = "\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest" + and ?IR2.0 = "\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest" + and ?IC1.0 = invar_ideal' and ?IC2.0=invar_real', + unfolded eI_defs \_adv_core_def \_usr_core_def, simplified]) + by (simp_all add: trace_eq_core fact1 fact2 fact3 fact4 ideal_s_rest'_def) +qed + + +subsubsection \Simplifying the final resource by moving the interfaces from core to rest\ + + +lemma connect[unfolded \_adv_core_def \_usr_core_def]: + fixes \_adv_restk \_adv_resta \_usr_restk \_usr_resta + defines "\ \ (\_adv_core \\<^sub>\ (\_adv_restk \\<^sub>\ \_adv_resta)) \\<^sub>\ (\_usr_core \\<^sub>\ (\_usr_restk \\<^sub>\ \_usr_resta))" + assumes [WT_intro]: "WT_rest \_adv_restk \_usr_restk I_key_rest key_rest" + and [WT_intro]: "WT_rest \_adv_resta \_usr_resta I_auth_rest auth_rest" + and "exception_\ \ \g D \" + shows "connect D (obsf_resource ideal_resource) = connect D (obsf_resource real_resource)" +proof - + note I_defs = \_adv_core_def \_usr_core_def + + have fact1: "\ \res RES (fused_resource.fuse ideal_core' ideal_rest') s \" + if "pred_prod I_key_rest I_auth_rest (snd (snd s))" "invar_ideal' (fst s)" + for s + unfolding assms(1) + apply(rule callee_invariant_on.WT_resource_of_oracle[where I="pred_prod invar_ideal' (\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest)"]) + subgoal by(rule fused_resource.callee_invariant_on_fuse)(rule WT_intro)+ + subgoal using that by(cases s)(simp) + done + + have fact2: "\ \res RES (fused_resource.fuse real_core' real_rest') s \" + if "pred_prod I_key_rest I_auth_rest (snd (snd s))" "invar_real' (fst s)" + for s + unfolding real_rest'_def assms(1) + apply(rule callee_invariant_on.WT_resource_of_oracle[where I="pred_prod invar_real' (\(_, s_rest). pred_prod I_key_rest I_auth_rest s_rest)"]) + subgoal by(rule fused_resource.callee_invariant_on_fuse)(rule WT_intro)+ + subgoal using that by(cases s)(simp) + done + + show ?thesis + unfolding attach_ideal attach_real + apply (rule connect_cong_trace[where \="exception_\ \"]) + apply (rule trace_eq_obsf_resourceI, subst trace_eq'_resource_of_oracle) + apply (rule trace_eq_sec[OF assms(2) assms(3)]) + subgoal by (rule assms(4)) + subgoal using WT_gpv_outs_gpv[OF assms(4)] by(simp add: I_defs assms(1) nempty_carrier) + subgoal using assms(2,3)[THEN WT_restD_rinit] by (intro WT_obsf_resource)(rule fact1; simp add: ideal_s_rest'_def) + subgoal using assms(2,3)[THEN WT_restD_rinit] by (intro WT_obsf_resource)(rule fact2; simp add: real_s_rest'_def ideal_s_rest'_def) + done +qed + +end + +end + + +end + + +subsection \Concrete security\ + +context one_time_pad begin + +lemma WT_enc_callee [WT_intro]: + "\_uniform (sec.Inp_Send ` carrier \) UNIV, \_uniform UNIV (key.Out_Alice ` carrier \) \\<^sub>\ \_uniform (sec.Inp_Send ` carrier \) UNIV \\<^sub>C CNV enc_callee () \" + by (rule WT_converter_of_callee) (auto 4 3 simp add: enc_callee_def stateless_callee_def image_def split!: key.ousr_alice.split) + +lemma WT_dec_callee [WT_intro]: + "\_uniform UNIV (sec.Out_Recv ` carrier \), \_uniform UNIV (key.Out_Bob ` carrier \) \\<^sub>\ \_uniform UNIV (sec.Out_Recv ` carrier \) \\<^sub>C CNV dec_callee () \" + by(rule WT_converter_of_callee)(auto simp add: dec_callee_def stateless_callee_def split!: sec.ousr_bob.splits) + +lemma pfinite_enc_callee [pfinite_intro]: + "pfinite_converter (\_uniform (sec.Inp_Send ` carrier \) UNIV) (\_uniform UNIV (key.Out_Alice ` carrier \) \\<^sub>\ \_uniform (sec.Inp_Send ` carrier \) UNIV) (CNV enc_callee ())" + apply(rule raw_converter_invariant.pfinite_converter_of_callee[where I="\_. True"]) + subgoal by unfold_locales(auto simp add: enc_callee_def stateless_callee_def) + subgoal by(auto simp add: enc_callee_def stateless_callee_def) + subgoal by simp + done + +lemma pfinite_dec_callee [pfinite_intro]: + "pfinite_converter (\_uniform UNIV (sec.Out_Recv ` carrier \)) (\_uniform UNIV (key.Out_Bob ` carrier \) \\<^sub>\ \_uniform UNIV (sec.Out_Recv ` carrier \)) (CNV dec_callee ())" + apply(rule raw_converter_invariant.pfinite_converter_of_callee[where I="\_. True"]) + subgoal by unfold_locales(auto simp add: dec_callee_def stateless_callee_def) + subgoal by(auto simp add: dec_callee_def stateless_callee_def) + subgoal by simp + done + +context + fixes + key_rest :: "('key_s_rest, key.event, 'key_iadv_rest, 'key_iusr_rest, 'key_oadv_rest, 'key_ousr_rest) rest_wstate" and + auth_rest :: "('auth_s_rest, auth.event, 'auth_iadv_rest, 'auth_iusr_rest, 'auth_oadv_rest, 'auth_ousr_rest) rest_wstate" and + \_adv_restk and \_adv_resta and \_usr_restk and \_usr_resta and I_key_rest and I_auth_rest + assumes + WT_key_rest [WT_intro]: "WT_rest \_adv_restk \_usr_restk I_key_rest key_rest" and + WT_auth_rest [WT_intro]: "WT_rest \_adv_resta \_usr_resta I_auth_rest auth_rest" +begin + +theorem secure: + defines "\_real \ ((\_full \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (sec.Inp_Fedit ` (carrier \)) UNIV))) \\<^sub>\ (\_adv_restk \\<^sub>\ \_adv_resta))" + and "\_common_core \ \_uniform (sec.Inp_Send ` (carrier \)) UNIV \\<^sub>\ \_uniform UNIV (sec.Out_Recv ` carrier \)" + and "\_common_rest \ \_usr_restk \\<^sub>\ \_usr_resta" + and "\_ideal \ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (sec.Inp_Fedit ` (carrier \)) UNIV)) \\<^sub>\ (\_adv_restk \\<^sub>\ \_adv_resta)" + shows "constructive_security_obsf (real_resource TYPE(_) TYPE(_) key_rest auth_rest) (sec.resource (ideal_rest key_rest auth_rest)) (sim |\<^sub>= 1\<^sub>C) \_real \_ideal (\_common_core \\<^sub>\ \_common_rest) \ 0" +proof + let ?\_common = "\_common_core \\<^sub>\ \_common_rest" + show "\_real \\<^sub>\ ?\_common \res real_resource TYPE(_) TYPE(_) key_rest auth_rest \" + unfolding \_real_def \_common_core_def \_common_rest_def real_resource_def attach_c1f22_c1f22_def wiring_c1r22_c1r22_def fused_wiring_def + by(rule WT_intro | simp )+ + + show [WT_intro]: "\_ideal \\<^sub>\ ?\_common \res sec.resource (ideal_rest key_rest auth_rest) \" + unfolding \_common_core_def \_common_rest_def \_ideal_def ideal_rest_def + by(rule WT_intro)+ simp + + show [WT_intro]: "\_real, \_ideal \\<^sub>C sim |\<^sub>= 1\<^sub>C \" + unfolding \_real_def \_ideal_def + apply(rule WT_intro)+ + subgoal + unfolding sim_def Let_def look_callee_def + apply (fold conv_callee_parallel_id_right[where s'="()"]) + apply (fold conv_callee_parallel_id_left[where s="()"]) + apply (subst ldummy_converter_of_callee) + apply (rule WT_converter_of_callee) + by (auto simp add: id_oracle_def map_gpv_conv_bind[symmetric] map_lift_spmf + split: auth.oadv_look.split option.split ) + by (rule WT_intro) + + show "pfinite_converter \_real \_ideal (sim |\<^sub>= 1\<^sub>C)" + unfolding \_real_def \_ideal_def + apply(rule pfinite_intro)+ + subgoal + unfolding sim_def Let_def look_callee_def + apply (fold conv_callee_parallel_id_right[where s'="()"]) + apply (fold conv_callee_parallel_id_left[where s="()"]) + apply (subst ldummy_converter_of_callee) + apply(rule raw_converter_invariant.pfinite_converter_of_callee[where I="\_. True"]) + subgoal + by unfold_locales (auto split!: sum.split sec.oadv_look.split option.split + simp add: left_gpv_map id_oracle_def intro!: WT_intro WT_gpv_right_gpv WT_gpv_left_gpv) + by (auto split!: sum.splits sec.oadv_look.splits option.splits) + by (rule pfinite_intro) + + assume WT [WT_intro]: "exception_\ (\_real \\<^sub>\ ?\_common) \g \ \" + show "advantage \ (obsf_resource ((sim |\<^sub>= 1\<^sub>C) |\<^sub>= (1\<^sub>C |\<^sub>= 1\<^sub>C) \ sec.resource (ideal_rest key_rest auth_rest))) + (obsf_resource (real_resource TYPE(_) TYPE(_) key_rest auth_rest)) \ 0" + using connect[OF WT_key_rest, OF WT_auth_rest, OF WT[unfolded assms(1, 2, 3)]] + unfolding advantage_def by (simp add: ideal_resource_def) +qed simp + +end + +end + + +subsection \Asymptotic security\ + +locale one_time_pad' = + fixes \ :: "security \ ('msg, 'more) boolean_algebra_scheme" + assumes one_time_pad [locale_witness]: "\\. one_time_pad (\ \)" +begin + +sublocale one_time_pad "\ \" for \ .. + +definition real_resource' where "real_resource' rest1 rest2 \ = real_resource TYPE(_) TYPE(_) \ (rest1 \) (rest2 \)" +definition ideal_resource' where "ideal_resource' rest1 rest2 \ = sec.resource \ (ideal_rest (rest1 \) (rest2 \))" +definition sim' where "sim' \ = (sim |\<^sub>= 1\<^sub>C)" + +context + fixes + key_rest :: "nat \ ('key_s_rest, key.event, 'key_iadv_rest, 'key_iusr_rest, 'key_oadv_rest, 'key_ousr_rest) rest_wstate" and + auth_rest :: "nat \ ('auth_s_rest, auth.event, 'auth_iadv_rest, 'auth_iusr_rest, 'auth_oadv_rest, 'auth_ousr_rest) rest_wstate" and + \_adv_restk and \_adv_resta and \_usr_restk and \_usr_resta and I_key_rest and I_auth_rest + assumes + WT_key_res: "\\. WT_rest (\_adv_restk \) (\_usr_restk \) (I_key_rest \) (key_rest \)" and + WT_auth_rest: "\\. WT_rest (\_adv_resta \) (\_usr_resta \) (I_auth_rest \) (auth_rest \)" +begin + +theorem secure': + defines "\_real \ \\. ((\_full \\<^sub>\ (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (sec.Inp_Fedit ` (carrier (\ \))) UNIV))) \\<^sub>\ (\_adv_restk \ \\<^sub>\ \_adv_resta \))" + and "\_common \ \\. ((\_uniform (sec.Inp_Send ` (carrier (\ \))) UNIV \\<^sub>\ \_uniform UNIV (sec.Out_Recv ` carrier (\ \))) \\<^sub>\ (\_usr_restk \ \\<^sub>\ \_usr_resta \))" + and "\_ideal \ \\. (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (sec.Inp_Fedit ` (carrier (\ \))) UNIV)) \\<^sub>\ (\_adv_restk \ \\<^sub>\ \_adv_resta \)" + shows "constructive_security_obsf' (real_resource' key_rest auth_rest) (ideal_resource' key_rest auth_rest) sim' \_real \_ideal \_common \" +proof(rule constructive_security_obsf'I) + show "constructive_security_obsf (real_resource' key_rest auth_rest \) (ideal_resource' key_rest auth_rest \) + (sim' \) (\_real \) (\_ideal \) (\_common \) (\ \) 0" for \ + unfolding real_resource'_def ideal_resource'_def sim'_def \_real_def \_common_def \_ideal_def + by(rule secure)(rule WT_key_res WT_auth_rest)+ +qed simp + +end + + +end + +end diff --git a/thys/Constructive_Cryptography_CM/Fold_Spmf.thy b/thys/Constructive_Cryptography_CM/Fold_Spmf.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Fold_Spmf.thy @@ -0,0 +1,200 @@ +theory Fold_Spmf + imports + More_CC +begin + +primrec (transfer) + foldl_spmf :: "('b \ 'a \ 'b spmf) \ 'b spmf \ 'a list \ 'b spmf" + where + foldl_spmf_Nil: "foldl_spmf f p [] = p" + | foldl_spmf_Cons: "foldl_spmf f p (x # xs) = foldl_spmf f (bind_spmf p (\a. f a x)) xs" + +lemma foldl_spmf_return_pmf_None [simp]: + "foldl_spmf f (return_pmf None) xs = return_pmf None" + by(induction xs) simp_all + +lemma foldl_spmf_bind_spmf: "foldl_spmf f (bind_spmf p g) xs = bind_spmf p (\a. foldl_spmf f (g a) xs)" + by(induction xs arbitrary: g) simp_all + +lemma bind_foldl_spmf_return: + "bind_spmf p (\x. foldl_spmf f (return_spmf x) xs) = foldl_spmf f p xs" + by(simp add: foldl_spmf_bind_spmf[symmetric]) + +lemma foldl_spmf_map [simp]: "foldl_spmf f p (map g xs) = foldl_spmf (map_fun id (map_fun g id) f) p xs" + by(induction xs arbitrary: p) simp_all + + +lemma foldl_spmf_identity [simp]: "foldl_spmf (\s x. return_spmf s) p xs = p" + by(induction xs arbitrary: p) simp_all + +lemma foldl_spmf_conv_foldl: + "foldl_spmf (\s x. return_spmf (f s x)) p xs = map_spmf (\s. foldl f s xs) p" + by(induction xs arbitrary: p)(simp_all add: map_spmf_conv_bind_spmf[symmetric] spmf.map_comp o_def) + +lemma foldl_spmf_Cons': + "foldl_spmf f (return_spmf a) (x # xs) = bind_spmf (f a x) (\a'. foldl_spmf f (return_spmf a') xs)" + by(simp add: bind_foldl_spmf_return) + +lemma foldl_spmf_append: "foldl_spmf f p (xs @ ys) = foldl_spmf f (foldl_spmf f p xs) ys" + by(induction xs arbitrary: p) simp_all + +lemma + foldl_spmf_helper: + assumes "\x. h (f x) = x" + assumes "\x. f (h x) = x" + shows "foldl_spmf (\a e. map_spmf h (g (f a) e)) acc es = + map_spmf h (foldl_spmf g (map_spmf f acc) es)" + using assms proof (induction es arbitrary: acc) + case (Cons a es) + then show ?case + by (simp add: spmf.map_comp map_bind_spmf bind_map_spmf o_def) +qed (simp add: map_spmf_conv_bind_spmf) + +lemma + foldl_spmf_helper2: + assumes "\x y. p (f x y) = x" + assumes "\x y. q (f x y) = y" + assumes "\x. f (p x) (q x) = x" + shows "foldl_spmf (\a e. map_spmf (f (p a)) (g (q a) e)) acc es = + bind_spmf acc (\acc'. map_spmf (f (p acc')) (foldl_spmf g (return_spmf (q acc')) es))" + proof (induction es arbitrary: acc) + note [simp] = spmf.map_comp map_bind_spmf bind_map_spmf o_def + case (Cons e es) + then show ?case + apply (simp add: map_spmf_conv_bind_spmf assms) + apply (subst bind_spmf_assoc[symmetric]) + by (simp add: bind_foldl_spmf_return) +qed (simp add: assms(3)) + +lemma foldl_pair_constl: "foldl (\s e. map_prod (\_. c) (\r. f r e) s) (c, sr) l = + Pair c (foldl (\s e. f s e) sr l)" + by (induction l arbitrary: sr) (auto simp add: map_prod_def split_def) + +lemma foldl_spmf_pair_left: + "foldl_spmf (\(l, r) e. map_spmf (\l'. (l', r)) (f l e)) (return_spmf (l, r)) es = + map_spmf (\l'. (l', r)) (foldl_spmf f (return_spmf l) es)" + apply (induction es arbitrary: l) + apply simp_all + apply (subst (2) map_spmf_conv_bind_spmf) + apply (subst foldl_spmf_bind_spmf) + apply (subst (2) bind_foldl_spmf_return[symmetric]) + by (simp add: map_spmf_conv_bind_spmf) + +lemma foldl_spmf_pair_left2: + "foldl_spmf (\(l, _) e. map_spmf (\l'. (l', c')) (f l e)) (return_spmf (l, c)) es = + map_spmf (\l'. (l', if es = [] then c else c')) (foldl_spmf f (return_spmf l) es)" + apply (induction es arbitrary: l c c') + apply simp_all + apply (subst (2) map_spmf_conv_bind_spmf) + apply (subst foldl_spmf_bind_spmf) + apply (subst (2) bind_foldl_spmf_return[symmetric]) + by (simp add: map_spmf_conv_bind_spmf) + +lemma foldl_pair_constr: "foldl (\s e. map_prod (\l. f l e) (\_. c) s) (sl, c) l = + Pair (foldl (\s e. f s e) sl l) c" + by (induction l arbitrary: sl) (auto simp add: map_prod_def split_def) + +lemma foldl_spmf_pair_right: + "foldl_spmf (\(l, r) e. map_spmf (\r'. (l, r')) (f r e)) (return_spmf (l, r)) es = + map_spmf (\r'. (l, r')) (foldl_spmf f (return_spmf r) es)" + apply (induction es arbitrary: r) + apply simp_all + apply (subst (2) map_spmf_conv_bind_spmf) + apply (subst foldl_spmf_bind_spmf) + apply (subst (2) bind_foldl_spmf_return[symmetric]) + by (simp add: map_spmf_conv_bind_spmf) + +lemma foldl_spmf_pair_right2: + "foldl_spmf (\(_, r) e. map_spmf (\r'. (c', r')) (f r e)) (return_spmf (c, r)) es = + map_spmf (\r'. (if es = [] then c else c', r')) (foldl_spmf f (return_spmf r) es)" + apply (induction es arbitrary: r c c') + apply simp_all + apply (subst (2) map_spmf_conv_bind_spmf) + apply (subst foldl_spmf_bind_spmf) + apply (subst (2) bind_foldl_spmf_return[symmetric]) + by (auto simp add: map_spmf_conv_bind_spmf split_def) + +lemma foldl_spmf_pair_right3: + "foldl_spmf (\(l, r) e. map_spmf (Pair (g e)) (f r e)) (return_spmf (l, r)) es = + map_spmf (Pair (if es = [] then l else g (last es))) (foldl_spmf f (return_spmf r) es)" + apply (induction es arbitrary: r l) + apply simp_all + apply (subst (2) map_spmf_conv_bind_spmf) + apply (subst foldl_spmf_bind_spmf) + apply (subst (2) bind_foldl_spmf_return[symmetric]) + by (clarsimp simp add: split_def map_bind_spmf o_def) + +lemma foldl_pullout: "bind_spmf f (\x. bind_spmf (foldl_spmf g init (events x)) (\y. h x y)) = + bind_spmf (bind_spmf f (\x. foldl_spmf (\(l, r) e. map_spmf (Pair l) (g r e)) (map_spmf (Pair x) init) (events x))) + (\(x, y). h x y)" for f g h init events + apply (simp add: foldl_spmf_helper2[where f=Pair and p=fst and q=snd, simplified] split_def) + apply (clarsimp simp add: map_spmf_conv_bind_spmf) + by (subst bind_spmf_assoc[symmetric]) (auto simp add: bind_foldl_spmf_return) + +lemma bind_foldl_spmf_pair_append: " + bind_spmf + (foldl_spmf (\(x, y) e. map_spmf (apfst ((@) x)) (f y e)) (return_spmf (a @ c, b)) es) + (\(x, y). g x y) = + bind_spmf + (foldl_spmf (\(x, y) e. map_spmf (apfst ((@) x)) (f y e)) (return_spmf (c, b)) es) + (\(x, y). g (a @ x) y)" + apply (induction es arbitrary: c b) + apply (simp_all add: split_def map_spmf_conv_bind_spmf apfst_def map_prod_def) + apply (subst (1 2) foldl_spmf_bind_spmf) + by simp + +lemma foldl_spmf_chain: " +(foldl_spmf (\(oevents, s_event) event. map_spmf (map_prod ((@) oevents) id) (fff s_event event)) (return_spmf ([], s_event)) ievents) + \ (\(oevents, s_event'). foldl_spmf ggg (return_spmf s_core) oevents + \ (\s_core'. return_spmf (f s_core' s_event'))) = +foldl_spmf (\(s_event, s_core) event. fff s_event event \ (\(oevents, s_event'). + map_spmf (Pair s_event') (foldl_spmf ggg (return_spmf s_core) oevents))) (return_spmf (s_event, s_core)) ievents + \ (\(s_event', s_core'). return_spmf (f s_core' s_event'))" +proof (induction ievents arbitrary: s_event s_core) + case Nil + show ?case by simp +next + case (Cons e es) + + show ?case + apply (subst (1 2) foldl_spmf_Cons') + apply (simp add: split_def) + apply (subst map_spmf_conv_bind_spmf) + apply simp + apply (rule bind_spmf_cong[OF refl]) + apply (subst (2) map_spmf_conv_bind_spmf) + apply simp + apply (subst Cons.IH[symmetric, simplified split_def]) + apply (subst bind_commute_spmf) + apply (subst (2) map_spmf_conv_bind_spmf[symmetric]) + apply (subst map_bind_spmf[symmetric, simplified o_def]) + apply (subst (1) foldl_spmf_bind_spmf[symmetric]) + apply (subst (3) map_spmf_conv_bind_spmf) + apply (simp add: foldl_spmf_append[symmetric] map_prod_def split_def) + subgoal for x + apply (cases x) + subgoal for a b + apply (simp add: split_def) + apply (subst bind_foldl_spmf_pair_append[where c="[]" and a=a and b=b and es=es, simplified apfst_def map_prod_def append_Nil2 split_def id_def]) + by simp + done + done +qed + + +\ \pauses\ +primrec pauses :: "'a list \ (unit, 'a, 'b) gpv" where + "pauses [] = Done ()" +| "pauses (x # xs) = Pause x (\_. pauses xs)" + +lemma WT_gpv_pauses [WT_intro]: + "\ \g pauses xs \" if "set xs \ outs_\ \" + using that by(induction xs) auto + +lemma exec_gpv_pauses: + "exec_gpv callee (pauses xs) s = + map_spmf (Pair ()) (foldl_spmf (map_fun id (map_fun id (map_spmf snd)) callee) (return_spmf s) xs)" + by(induction xs arbitrary: s)(simp_all add: split_def foldl_spmf_Cons' map_bind_spmf bind_map_spmf o_def del: foldl_spmf_Cons) + + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Fused_Resource.thy b/thys/Constructive_Cryptography_CM/Fused_Resource.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Fused_Resource.thy @@ -0,0 +1,1721 @@ +theory Fused_Resource imports + Fold_Spmf +begin + +(* TODO: move these somewhere *) +context includes \.lifting begin +lift_definition e\ :: "('a, 'b) \ \ ('a, 'b \ 'c) \" is "\\ x. \ x \ UNIV" . + +lemma outs_\_e\[simp]: "outs_\ (e\ \) = outs_\ \" + by transfer simp + +lemma responses_\_e\ [simp]: "responses_\ (e\ \) x = responses_\ \ x \ UNIV" + by transfer simp + +lemma e\_map_\: "e\ (map_\ f g \) = map_\ f (apfst g) (e\ \)" + by transfer(auto simp add: fun_eq_iff intro: rev_image_eqI) + +lemma e\_inverse [simp]: "map_\ id fst (e\ \) = \" + by transfer auto +end +lifting_update \.lifting +lifting_forget \.lifting + + +section \Fused Resource\ + +subsection \Event Oracles -- they generate events\ + +type_synonym + ('state, 'event, 'input, 'output) eoracle = "('state, 'input, 'output \ 'event list) oracle'" + +definition + parallel_eoracle :: " + ('s1, 'e1, 'i1, 'o1) eoracle \ ('s2, 'e2, 'i2, 'o2) eoracle \ + ('s1 \ 's2, 'e1 + 'e2, 'i1 + 'i2, 'o1 + 'o2) eoracle" + where + "parallel_eoracle eoracle1 eoracle2 state \ + comp + (map_spmf + (map_prod + (case_sum + (map_prod Inl (map Inl)) + (map_prod Inr (map Inr))) + id)) + (parallel_oracle eoracle1 eoracle2 state)" + +definition + plus_eoracle :: " + ('s, 'e1, 'i1, 'o1) eoracle \ ('s, 'e2, 'i2, 'o2) eoracle \ + ('s, 'e1 + 'e2, 'i1 + 'i2, 'o1 + 'o2) eoracle" + where + "plus_eoracle eoracle1 eoracle2 state \ + comp + (map_spmf + (map_prod + (case_sum + (map_prod Inl (map Inl)) + (map_prod Inr (map Inr))) + id)) + (plus_oracle eoracle1 eoracle2 state)" + +definition + translate_eoracle :: " + ('s_event, 'e1, 'e2 list) oracle' \ ('s_event \ 's, 'e1, 'i, 'o) eoracle \ + ('s_event \ 's, 'e2, 'i, 'o) eoracle" + where + "translate_eoracle translator eoracle state inp \ + bind_spmf + (eoracle state inp) + (\((out, e_in), s). + let conc = (\(es, st) e. map_spmf (map_prod ((@) es) id) (translator st e)) in do { + (e_out, s_event) \ foldl_spmf conc (return_spmf ([], fst s)) e_in; + return_spmf ((out, e_out), s_event, snd s) + })" + + +subsection \Event Handlers -- they absorb (and silently handle) events\ + +type_synonym + ('state, 'event) handler = "'state \ 'event \ 'state spmf" + +fun + parallel_handler :: "('s1, 'e1) handler \ ('s2, 'e2) handler \ ('s1 \ 's2, 'e1 + 'e2) handler" + where + "parallel_handler left _ s (Inl e1) = map_spmf (\s1'. (s1', snd s)) (left (fst s) e1)" + | "parallel_handler _ right s (Inr e2) = map_spmf (\s2'. (fst s, s2')) (right (snd s) e2)" + +definition + plus_handler :: "('s, 'e1) handler \ ('s, 'e2) handler \ ('s, 'e1 + 'e2) handler" + where + "plus_handler left right s \ case_sum (left s) (right s)" + +lemma parallel_handler_left: + "map_fun id (map_fun Inl id) (parallel_handler left right) = + (\(s_l, s_r) q. map_spmf (\s_l'. (s_l', s_r)) (left s_l q))" + by force + +lemma parallel_handler_right: + "map_fun id (map_fun Inr id) (parallel_handler left right) = + (\(s_l, s_r) q. map_spmf (\s_r'. (s_l, s_r')) (right s_r q))" + by force + +lemma in_set_spmf_parallel_handler: + "s' \ set_spmf (parallel_handler left right s x) \ + (case x of Inl e \ fst s' \ set_spmf (left (fst s) e) \ snd s' = snd s + | Inr e \ snd s' \ set_spmf (right (snd s) e) \ fst s' = fst s)" + by(cases s; cases s'; auto split: sum.split) + +subsection \Fused Resource Construction\ + +codatatype + ('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core = + Core + (cpoke: "('s_core, 'event) handler") + (cfunc_adv: "('s_core, 'iadv_core, 'oadv_core) oracle'") + (cfunc_usr: "('s_core, 'iusr_core, 'ousr_core) oracle'") + +declare core.sel_transfer[transfer_rule del] +declare core.ctr_transfer[transfer_rule del] +declare core.case_transfer[transfer_rule del] + +context + includes lifting_syntax +begin + +inductive + rel_core':: " + ('s_core \ 's_core' \ bool) \ + ('event \ 'event' \ bool) \ + ('iadv_core \ 'iadv_core' \ bool) \ + ('iusr_core \ 'iusr_core' \ bool) \ + ('oadv_core \ 'oadv_core' \ bool) \ + ('ousr_core \ 'ousr_core' \ bool) \ + ('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core \ + ('s_core', 'event', 'iadv_core', 'iusr_core', 'oadv_core', 'ousr_core') core \ bool" + for S E IA IU OA OU + where "rel_core' S E IA IU OA OU (Core cpoke cfunc_adv cfunc_usr) (Core cpoke' cfunc_adv' cfunc_usr')" + if + "(S ===> E ===> rel_spmf S) cpoke cpoke'" and + "(S ===> IA ===> rel_spmf (rel_prod OA S)) cfunc_adv cfunc_adv'" and + "(S ===> IU ===> rel_spmf (rel_prod OU S)) cfunc_usr cfunc_usr'" + for cpoke cfunc_adv cfunc_usr + +inductive_simps + rel_core'_simps [simp]: + "rel_core' S E IA IU OA OU (Core cpoke' cfunc_adv' cfunc_usr') (Core cpoke'' cfunc_adv'' cfunc_usr'')" + +lemma + rel_core'_eq [relator_eq]: + "rel_core' (=) (=) (=) (=) (=) (=) = (=)" + apply(intro ext) + subgoal for x y by(cases x; cases y)(auto simp add: fun_eq_iff rel_fun_def relator_eq) + done + +lemma + rel_core'_mono [relator_mono]: + "rel_core' S E IA IU OA OU \ rel_core' S E' IA' IU' OA' OU'" + if "E' \ E" "IA' \ IA" "IU' \ IU" "OA \ OA'" "OU \ OU'" + apply(rule predicate2I) + subgoal for x y + apply(cases x; cases y; clarsimp; intro conjI) + apply(erule rel_fun_mono rel_spmf_mono prod.rel_mono[THEN predicate2D, rotated -1] | + rule impI that order_refl | erule that[THEN predicate2D] | erule rel_spmf_mono | assumption)+ + done + done + +lemma + cpoke_parametric [transfer_rule]: + "(rel_core' S E IA IU OA OU ===> S ===> E ===> rel_spmf S) cpoke cpoke" + by(rule rel_funI; erule rel_core'.cases; simp) + +lemma + cfunc_adv_parametric [transfer_rule]: + "(rel_core' S E IA IU OA OU ===> S ===> IA ===> rel_spmf (rel_prod OA S)) cfunc_adv cfunc_adv" + by(rule rel_funI; erule rel_core'.cases; simp) + +lemma + cfunc_usr_parametric [transfer_rule]: + "(rel_core' S E IA IU OA OU ===> S ===> IU ===> rel_spmf (rel_prod OU S)) cfunc_usr cfunc_usr" + by(rule rel_funI; erule rel_core'.cases; simp) + +lemma + Core_parametric [transfer_rule]: + "((S ===> E ===> rel_spmf S) ===> (S ===> IA ===> rel_spmf (rel_prod OA S)) ===> (S ===> IU ===> rel_spmf (rel_prod OU S)) + ===> rel_core' S E IA IU OA OU) Core Core" + by(rule rel_funI)+ simp + +lemma + case_core_parametric [transfer_rule]: + "(((S ===> E ===> rel_spmf S) ===> + (S ===> IA ===> rel_spmf (rel_prod OA S)) ===> + (S ===> IU ===> rel_spmf (rel_prod OU S)) ===> X) ===> + rel_core' S E IA IU OA OU ===> X) case_core case_core" + by(rule rel_funI)+(auto 4 4 split: core.split dest: rel_funD) + +lemma + corec_core_parametric [transfer_rule]: + "((X ===> S ===> E ===> rel_spmf S) ===> + (X ===> S ===> IA ===> rel_spmf (rel_prod OA S)) ===> + (X ===> S ===> IU ===> rel_spmf (rel_prod OU S)) ===> + X ===> rel_core' S E IA IU OA OU) corec_core corec_core" + by(rule rel_funI)+(auto simp add: core.corec dest: rel_funD) + +primcorec map_core' :: + "('event' \ 'event) \ + ('iadv_core' \ 'iadv_core) \ + ('iusr_core' \ 'iusr_core) \ + ('oadv_core \ 'oadv_core') \ + ('ousr_core \ 'ousr_core') \ + ('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core \ + ('s_core, 'event', 'iadv_core', 'iusr_core', 'oadv_core', 'ousr_core') core" + where + "cpoke (map_core' e ia iu oa ou core) = (id ---> e ---> id) (cpoke core)" +| "cfunc_adv (map_core' e ia iu oa ou core) = (id ---> ia ---> map_spmf (map_prod oa id)) (cfunc_adv core)" +| "cfunc_usr (map_core' e ia iu oa ou core) = (id ---> iu ---> map_spmf (map_prod ou id)) (cfunc_usr core)" + +lemmas map_core'_simps [simp] = map_core'.ctr[where core="Core _ _ _", simplified] + +parametric_constant map_core'_parametric[transfer_rule]: map_core'_def + +lemma core'_rel_Grp: + "rel_core' (=) (BNF_Def.Grp UNIV e)\\ (BNF_Def.Grp UNIV ia)\\ (BNF_Def.Grp UNIV iu)\\ (BNF_Def.Grp UNIV oa) (BNF_Def.Grp UNIV ou) + = BNF_Def.Grp UNIV (map_core' e ia iu oa ou)" + apply(intro ext) + subgoal for x y + apply(cases x; cases y; clarsimp) + apply(subst (2 4 6) eq_alt_conversep) + apply(subst (2 3 4) eq_alt) + apply(simp add: pmf.rel_Grp option.rel_Grp prod.rel_Grp rel_fun_conversep_grp_grp) + apply(auto simp add: Grp_def spmf.map_id[abs_def] id_def[symmetric]) + done + done + +end + +inductive WT_core :: "('iadv, 'oadv) \ \ ('iusr, 'ousr) \ \ ('s \ bool) \ ('s, 'event, 'iadv, 'iusr, 'oadv, 'ousr) core \ bool" + for \_adv \_usr I core where + "WT_core \_adv \_usr I core" if + "\s e s'. \ s' \ set_spmf (cpoke core s e); I s \ \ I s'" + "\s x y s'. \ (y, s') \ set_spmf (cfunc_adv core s x); x \ outs_\ \_adv; I s \ \ y \ responses_\ \_adv x \ I s'" + "\s x y s'. \ (y, s') \ set_spmf (cfunc_usr core s x); x \ outs_\ \_usr; I s \ \ y \ responses_\ \_usr x \ I s'" + +lemma WT_coreD: + assumes "WT_core \_adv \_usr I core" + shows WT_coreD_cpoke: "\s e s'. \ s' \ set_spmf (cpoke core s e); I s \ \ I s'" + and WT_coreD_cfunc_adv: "\s x y s'. \ (y, s') \ set_spmf (cfunc_adv core s x); x \ outs_\ \_adv; I s \ \ y \ responses_\ \_adv x \ I s'" + and WT_coreD_cfund_usr: "\s x y s'. \ (y, s') \ set_spmf (cfunc_usr core s x); x \ outs_\ \_usr; I s \ \ y \ responses_\ \_usr x \ I s'" + using assms by(auto elim!: WT_core.cases) + +lemma WT_coreD_foldl_spmf_cpoke: + assumes "WT_core \_adv \_usr I core" + and "s' \ set_spmf (foldl_spmf (cpoke core) p es)" + and "\s \ set_spmf p. I s" + shows "I s'" + using assms(2, 3) + by(induction es arbitrary: p)(fastforce dest: WT_coreD_cpoke[OF assms(1)] simp add: bind_UNION)+ + +lemma WT_core_trivial: + assumes adv: "\s. \_adv \c cfunc_adv core s \" + and usr: "\s. \_usr \c cfunc_usr core s \" + shows "WT_core \_adv \_usr (\_. True) core" + by(rule WT_core.intros)(auto dest: assms[THEN WT_calleeD]) + +codatatype + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme = + Rest + (rinit: "'more") + (rfunc_adv: "('s_rest, 'event, 'iadv_rest, 'oadv_rest) eoracle") + (rfunc_usr: "('s_rest, 'event, 'iusr_rest, 'ousr_rest) eoracle") + +declare rest_scheme.sel_transfer[transfer_rule del] +declare rest_scheme.ctr_transfer[transfer_rule del] +declare rest_scheme.case_transfer[transfer_rule del] + +context + includes lifting_syntax +begin + +inductive + rel_rest':: " + ('s_rest \ 's_rest' \ bool) \ + ('event \ 'event' \ bool) \ + ('iadv_rest \ 'iadv_rest' \ bool) \ + ('iusr_rest \ 'iusr_rest' \ bool) \ + ('oadv_rest \ 'oadv_rest' \ bool) \ + ('ousr_rest \ 'ousr_rest' \ bool) \ + ('more \ 'more' \ bool) \ + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme \ + ('s_rest', 'event', 'iadv_rest', 'iusr_rest', 'oadv_rest', 'ousr_rest', 'more') rest_scheme \ bool" + for S E IA IU OA OU M + where "rel_rest' S E IA IU OA OU M (Rest rinit rfunc_adv rfunc_usr) (Rest rinit' rfunc_adv' rfunc_usr')" + if + "M rinit rinit'" and + "(S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) rfunc_adv rfunc_adv'" and + "(S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) rfunc_usr rfunc_usr'" + for rinit rfunc_adv rfunc_usr + +inductive_simps + rel_rest'_simps [simp]: + "rel_rest' S E IA IU OA OU M (Rest rinit' rfunc_adv' rfunc_usr') (Rest rinit'' rfunc_adv'' rfunc_usr'')" + +lemma + rel_rest'_eq [relator_eq]: + "rel_rest' (=) (=) (=) (=) (=) (=) (=) = (=)" + apply(intro ext) + subgoal for x y by(cases x; cases y)(auto simp add: fun_eq_iff rel_fun_def relator_eq) + done + +lemma + rel_rest'_mono [relator_mono]: + "rel_rest' S E IA IU OA OU M \ rel_rest' S E' IA' IU' OA' OU' M'" + if "E \ E'" "IA' \ IA" "IU' \ IU" "OA \ OA'" "OU \ OU'" "M \ M'" + apply(rule predicate2I) + subgoal for x y + apply(cases x; cases y; clarsimp; intro conjI) + apply(erule rel_fun_mono rel_spmf_mono prod.rel_mono[THEN predicate2D, rotated -1] | + rule impI that order_refl prod.rel_mono list.rel_mono | erule that[THEN predicate2D] | erule rel_spmf_mono | assumption)+ + done + done + +lemma rel_rest'_sel: "rel_rest' S E IA IU OA OU M rest1 rest2" + if "M (rinit rest1) (rinit rest2)" + and "(S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) (rfunc_adv rest1) (rfunc_adv rest2)" + and "(S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) (rfunc_usr rest1) (rfunc_usr rest2)" + using that by(cases rest1; cases rest2) simp + +lemma rinit_parametric [transfer_rule]: "(rel_rest' S E IA IU OA OU M ===> M) rinit rinit" + by(rule rel_funI; erule rel_rest'.cases; simp) + +lemma rfunc_adv_parametric [transfer_rule]: + "(rel_rest' S E IA IU OA OU M ===> S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) rfunc_adv rfunc_adv" + by(rule rel_funI; erule rel_rest'.cases; simp) + +lemma rfunc_usr_parametric [transfer_rule]: + "(rel_rest' S E IA IU OA OU M ===> S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) rfunc_usr rfunc_usr" + by(rule rel_funI; erule rel_rest'.cases; simp) + +lemma Rest_parametric [transfer_rule]: + "(M ===> (S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) + ===> (S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) + ===> rel_rest' S E IA IU OA OU M) Rest Rest" + by(rule rel_funI)+ simp + +lemma case_rest_scheme_parametric [transfer_rule]: + "((M ===> + (S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) ===> + (S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) ===> X) ===> + rel_rest' S E IA IU OA OU M ===> X) case_rest_scheme case_rest_scheme" + by(rule rel_funI)+(auto 4 4 split: rest_scheme.split dest: rel_funD) + +lemma corec_rest_scheme_parametric [transfer_rule]: + "((X ===> M) ===> + (X ===> S ===> IA ===> rel_spmf (rel_prod (rel_prod OA (list_all2 E)) S)) ===> + (X ===> S ===> IU ===> rel_spmf (rel_prod (rel_prod OU (list_all2 E)) S)) ===> + X ===> rel_rest' S E IA IU OA OU M) corec_rest_scheme corec_rest_scheme" + by(rule rel_funI)+(auto simp add: rest_scheme.corec dest: rel_funD) + +primcorec map_rest' :: + "('event \ 'event') \ + ('iadv_rest' \ 'iadv_rest) \ + ('iusr_rest' \ 'iusr_rest) \ + ('oadv_rest \ 'oadv_rest') \ + ('ousr_rest \ 'ousr_rest') \ + ('more \ 'more') \ + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme \ + ('s_rest, 'event', 'iadv_rest', 'iusr_rest', 'oadv_rest', 'ousr_rest', 'more') rest_scheme" + where + "rinit (map_rest' e ia iu oa ou m rest) = m (rinit rest)" +| "rfunc_adv (map_rest' e ia iu oa ou m rest) = + (id ---> ia ---> map_spmf (map_prod (map_prod oa (map e)) id)) (rfunc_adv rest)" +| "rfunc_usr (map_rest' e ia iu oa ou m rest) = + (id ---> iu ---> map_spmf (map_prod (map_prod ou (map e)) id)) (rfunc_usr rest)" + +lemmas map_rest'_simps [simp] = map_rest'.ctr[where rest="Rest _ _ _", simplified] + +parametric_constant map_rest'_parametric[transfer_rule]: map_rest'_def + +lemma rest'_rel_Grp: + "rel_rest' (=) (BNF_Def.Grp UNIV e) (BNF_Def.Grp UNIV ia)\\ (BNF_Def.Grp UNIV iu)\\ (BNF_Def.Grp UNIV oa) (BNF_Def.Grp UNIV ou) (BNF_Def.Grp UNIV m) + = BNF_Def.Grp UNIV (map_rest' e ia iu oa ou m)" + apply(intro ext) + subgoal for x y + apply(cases x; cases y; clarsimp) + apply(subst (2 4) eq_alt_conversep) + apply(subst (2 3) eq_alt) + apply(simp add: pmf.rel_Grp list.rel_Grp option.rel_Grp prod.rel_Grp rel_fun_conversep_grp_grp) + apply(auto simp add: Grp_def spmf.map_id[abs_def] id_def[symmetric]) + done + done + +end + +type_synonym + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) rest_wstate = + "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 's_rest) rest_scheme" + +inductive WT_rest :: "('iadv, 'oadv) \ \ ('iusr, 'ousr) \ \ ('s \ bool) \ ('s, 'event, 'iadv, 'iusr, 'oadv, 'ousr) rest_wstate \ bool" + for \_adv \_usr I rest where + "WT_rest \_adv \_usr I rest" if + "\s x y es s'. \ ((y, es), s') \ set_spmf (rfunc_adv rest s x); x \ outs_\ \_adv; I s \ \ y \ responses_\ \_adv x \ I s'" + "\s x y es s'. \ ((y, es), s') \ set_spmf (rfunc_usr rest s x); x \ outs_\ \_usr; I s \ \ y \ responses_\ \_usr x \ I s'" + "I (rinit rest)" + +lemma WT_restD: + assumes "WT_rest \_adv \_usr I rest" + shows WT_restD_rfunc_adv: "\s x y es s'. \ ((y, es), s') \ set_spmf (rfunc_adv rest s x); x \ outs_\ \_adv; I s \ \ y \ responses_\ \_adv x \ I s'" + and WT_restD_rfunc_usr: "\s x y es s'. \ ((y, es), s') \ set_spmf (rfunc_usr rest s x); x \ outs_\ \_usr; I s \ \ y \ responses_\ \_usr x \ I s'" + and WT_restD_rinit: "I (rinit rest)" + using assms by(auto elim!: WT_rest.cases) + +abbreviation + fuse_cfunc :: " + ('o \ 'x) \ ('s_core, 'i, 'o) oracle' \ ('s_core \ 's_rest, 'i , 'x) oracle'" + where + "fuse_cfunc redirect cfunc state inp \ do { + let handle = map_prod redirect (prod.swap o Pair (snd state)); + (os_cfunc :: 'o \ 's_core) \ cfunc (fst state) inp; + return_spmf (handle os_cfunc) + }" + +abbreviation + fuse_rfunc :: " + ('o \ 'x) \ ('s_rest, 'e, 'i, 'o) eoracle \ ('s_core, 'e) handler \ + ('s_core \ 's_rest, 'i , 'x) oracle'" + where + "fuse_rfunc redirect rfunc notify state inp \ + bind_spmf + (rfunc (snd state) inp) + (\((o_rfunc, e_lst), s_rfunc). + bind_spmf + (foldl_spmf notify (return_spmf (fst state)) e_lst) + (\s_notify. return_spmf (redirect o_rfunc, s_notify, s_rfunc))) + " + + +locale fused_resource = + fixes + core :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" and + core_init :: 's_core +begin + +fun + fuse :: " + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'm) rest_scheme \ + ('s_core \ 's_rest, + ('iadv_core + 'iadv_rest) + ('iusr_core + 'iusr_rest), + ('oadv_core + 'oadv_rest) + ('ousr_core + 'ousr_rest)) oracle'" + where + "fuse rest state (Inl (Inl iadv_core)) = + fuse_cfunc (Inl o Inl) (cfunc_adv core) state iadv_core" + | "fuse rest state (Inl (Inr iadv_rest)) = + fuse_rfunc (Inl o Inr) (rfunc_adv rest) (cpoke core) state iadv_rest" + | "fuse rest state (Inr (Inl iusr_core)) = + fuse_cfunc (Inr o Inl) (cfunc_usr core) state iusr_core" + | "fuse rest state (Inr (Inr iusr_rest)) = + fuse_rfunc (Inr o Inr) (rfunc_usr rest) (cpoke core) state iusr_rest" + +case_of_simps fuse_case: fused_resource.fuse.simps + +lemma callee_invariant_on_fuse: + assumes "WT_core \_adv_core \_usr_core I_core core" + and "WT_rest \_adv_rest \_usr_rest I_rest rest" + shows "callee_invariant_on (fuse rest) (pred_prod I_core I_rest) ((\_adv_core \\<^sub>\ \_adv_rest) \\<^sub>\ (\_usr_core \\<^sub>\ \_usr_rest))" +proof(unfold_locales, goal_cases) + case (1 s x y s') + then show ?case using assms + by(cases s; cases s')(auto 4 4 dest: WT_restD WT_coreD WT_coreD_foldl_spmf_cpoke) +next + case (2 s) + show ?case + apply(rule WT_calleeI) + subgoal for x y s' using 2 assms + by (cases "(rest, s, x)" rule: fuse.cases) (auto simp add: pred_prod_beta dest: WT_restD WT_coreD ) + done +qed + +definition + resource :: " + ('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) rest_wstate \ + (('iadv_core + 'iadv_rest) + ('iusr_core + 'iusr_rest), + ('oadv_core + 'oadv_rest) + ('ousr_core + 'ousr_rest)) resource" + where + "resource rest = resource_of_oracle (fuse rest) (core_init, rinit rest)" + +lemma WT_resource [WT_intro]: + assumes "WT_core \_adv_core \_usr_core I_core core" + and "WT_rest \_adv_rest \_usr_rest I_rest rest" + and "I_core core_init" + shows "(\_adv_core \\<^sub>\ \_adv_rest) \\<^sub>\ (\_usr_core \\<^sub>\ \_usr_rest) \res resource rest \" +proof - + interpret callee_invariant_on "fuse rest" "pred_prod I_core I_rest" "(\_adv_core \\<^sub>\ \_adv_rest) \\<^sub>\ (\_usr_core \\<^sub>\ \_usr_rest)" + using assms(1,2) by(rule callee_invariant_on_fuse) + show ?thesis unfolding resource_def + by(rule WT_resource_of_oracle)(simp add: assms(3) WT_restD_rinit[OF assms(2)]) +qed + +end + +parametric_constant + fuse_parametric [transfer_rule]: fused_resource.fuse_case + +subsection \More helpful construction functions\ + +context + fixes + core1 :: "('s_core1, 'event1, 'iadv_core1, 'iusr_core1, 'oadv_core1, 'ousr_core1) core" and + core2 :: "('s_core2, 'event2, 'iadv_core2, 'iusr_core2, 'oadv_core2, 'ousr_core2) core" +begin + +primcorec parallel_core :: " + ('s_core1 \ 's_core2, 'event1 + 'event2, + 'iadv_core1 + 'iadv_core2, 'iusr_core1 + 'iusr_core2, + 'oadv_core1 + 'oadv_core2, 'ousr_core1 + 'ousr_core2) core" + where + "cpoke parallel_core = parallel_handler (cpoke core1) (cpoke core2)" + | "cfunc_adv parallel_core = parallel_oracle (cfunc_adv core1) (cfunc_adv core2)" + | "cfunc_usr parallel_core = parallel_oracle (cfunc_usr core1) (cfunc_usr core2)" + +end + + +context + fixes + cnv_adv :: "'s_adv \ 'iadv \ ('oadv \ 's_adv, 'iadv_core, 'oadv_core) gpv" and + cnv_usr :: "'s_usr \ 'iusr \ ('ousr \ 's_usr, 'iusr_core, 'ousr_core) gpv" and + core :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" +begin + +primcorec + attach_core :: "(('s_adv \ 's_usr) \ 's_core, 'event, 'iadv, 'iusr, 'oadv, 'ousr) core" + where + "cpoke attach_core = (\(s_advusr, s_core) event. + map_spmf (\s_core'. (s_advusr, s_core')) (cpoke core s_core event))" + | "cfunc_adv attach_core = (\((s_adv, s_usr), s_core) iadv. + map_spmf + (\((oadv, s_adv'), s_core'). (oadv, ((s_adv', s_usr), s_core'))) + (exec_gpv (cfunc_adv core) (cnv_adv s_adv iadv) s_core))" + | "cfunc_usr attach_core = (\((s_adv, s_usr), s_core) iusr. + map_spmf + (\((ousr, s_usr'), s_core'). (ousr, ((s_adv, s_usr'), s_core'))) + (exec_gpv (cfunc_usr core) (cnv_usr s_usr iusr) s_core))" + +end + + +lemma + attach_core_id_oracle_adv: "cfunc_adv (attach_core 1\<^sub>I cnv core) = + (\(s_cnv, s_core) q. map_spmf (\(out, s_core'). (out, s_cnv, s_core')) (cfunc_adv core s_core q))" + by(simp add: id_oracle_def split_def map_spmf_conv_bind_spmf) + +lemma + attach_core_id_oracle_usr: "cfunc_usr (attach_core cnv 1\<^sub>I core) = + (\(s_cnv, s_core) q. map_spmf (\(out, s_core'). (out, s_cnv, s_core')) (cfunc_usr core s_core q))" + by(simp add: id_oracle_def split_def map_spmf_conv_bind_spmf) + + +context + fixes + rest1 :: "('s_rest1, 'event1, 'iadv_rest1, 'iusr_rest1, 'oadv_rest1, 'ousr_rest1, 'more1) rest_scheme" and + rest2 :: "('s_rest2, 'event2, 'iadv_rest2, 'iusr_rest2, 'oadv_rest2, 'ousr_rest2, 'more2) rest_scheme" +begin + +primcorec parallel_rest :: " + ('s_rest1 \ 's_rest2, 'event1 + 'event2, 'iadv_rest1 + 'iadv_rest2, 'iusr_rest1 + 'iusr_rest2, + 'oadv_rest1 + 'oadv_rest2, 'ousr_rest1 + 'ousr_rest2, 'more1 \ 'more2) rest_scheme" + where + "rinit parallel_rest = (rinit rest1, rinit rest2)" + | "rfunc_adv parallel_rest = parallel_eoracle (rfunc_adv rest1) (rfunc_adv rest2)" + | "rfunc_usr parallel_rest = parallel_eoracle (rfunc_usr rest1) (rfunc_usr rest2)" + +end + +lemma WT_parallel_rest [WT_intro]: + "WT_rest (\_adv1 \\<^sub>\ \_adv2) (\_usr1 \\<^sub>\ \_usr2) (pred_prod I1 I2) (parallel_rest rest1 rest2)" + if "WT_rest \_adv1 \_usr1 I1 rest1" + and "WT_rest \_adv2 \_usr2 I2 rest2" + by(rule WT_rest.intros) + (auto 4 3 simp add: parallel_eoracle_def simp add: that[THEN WT_restD_rinit] dest: that[THEN WT_restD_rfunc_adv] that[THEN WT_restD_rfunc_usr]) + +context + fixes + cnv_adv :: "'s_adv \ 'iadv \ ('oadv \ 's_adv, 'iadv_rest, 'oadv_rest) gpv" and + cnv_usr :: "'s_usr \ 'iusr \ ('ousr \ 's_usr, 'iusr_rest, 'ousr_rest) gpv" and + f_init :: "'more \ 'more'" and + rest :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" +begin + +primcorec + attach_rest :: " + (('s_adv \ 's_usr) \ 's_rest, 'event, 'iadv, 'iusr, 'oadv, 'ousr, 'more') rest_scheme" + where + "rinit attach_rest = f_init (rinit rest)" + | "rfunc_adv attach_rest = (\((s_adv, s_usr), s_rest) iadv. + let orc_of = \orc (s, es) q. map_spmf (\ ((out, e), s'). (out, s', es @ e)) (orc s q) in + let eorc_of = \((oadv, s_adv'), (s_rest', es)). ((oadv, es), ((s_adv', s_usr), s_rest')) in + map_spmf eorc_of (exec_gpv (orc_of (rfunc_adv rest)) (cnv_adv s_adv iadv) (s_rest, [])))" + | "rfunc_usr attach_rest = (\((s_adv, s_usr), s_rest) iusr. + let orc_of = \orc (s, es) q. map_spmf (\ ((out, e), s'). (out, s', es @ e)) (orc s q) in + let eorc_of = \((ousr, s_usr'), (s_rest', es)). ((ousr, es), ((s_adv, s_usr'), s_rest')) in + map_spmf eorc_of (exec_gpv (orc_of (rfunc_usr rest)) (cnv_usr s_usr iusr) (s_rest, [])))" + +end + + +lemma + attach_rest_id_oracle_adv: "rfunc_adv (attach_rest 1\<^sub>I cnv f_init rest) = + (\(s_cnv, s_core) q. map_spmf (\(out, s_core'). (out, s_cnv, s_core')) (rfunc_adv rest s_core q))" + by(simp add: id_oracle_def split_def map_spmf_conv_bind_spmf fun_eq_iff) + +lemma + attach_rest_id_oracle_usr: "rfunc_usr (attach_rest cnv 1\<^sub>I f_init rest) = + (\(s_cnv, s_core) q. map_spmf (\(out, s_core'). (out, s_cnv, s_core')) (rfunc_usr rest s_core q))" + by(simp add: id_oracle_def split_def map_spmf_conv_bind_spmf) + + + +section \Traces\ + +type_synonym ('event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) trace_core = + "('event + 'iadv_core \ 'oadv_core + 'iusr_core \ 'ousr_core) list + \ ('event \ real) + \ ('iadv_core \ 'oadv_core spmf) + \ ('iusr_core \ 'ousr_core spmf)" + +context + fixes core :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" +begin + +primrec trace_core' :: "'s_core spmf \ ('event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) trace_core" where + "trace_core' S [] = + (\e. weight_spmf' (bind_spmf S (\s. cpoke core s e)), + \ia. bind_spmf S (\s. map_spmf fst (cfunc_adv core s ia)), + \iu. bind_spmf S (\s. map_spmf fst (cfunc_usr core s iu)))" +| "trace_core' S (obs # tr) = (case obs of + Inl e \ trace_core' (mk_lossless (bind_spmf S (\s. cpoke core s e))) tr + | Inr (Inl (ia, oa)) \ trace_core' (cond_spmf_fst (bind_spmf S (\s. cfunc_adv core s ia)) oa) tr + | Inr (Inr (iu, ou)) \ trace_core' (cond_spmf_fst (bind_spmf S (\s. cfunc_usr core s iu)) ou) tr + )" + +end + +declare trace_core'.simps [simp del] +case_of_simps trace_core'_unfold: trace_core'.simps[unfolded weight_spmf'_def] +simps_of_case trace_core'_simps [simp]: trace_core'_unfold + +context includes lifting_syntax begin + +lemma trace_core'_parametric [transfer_rule]: + "(rel_core' S E IA IU (=) (=) ===> + rel_spmf S ===> + list_all2 (rel_sum E (rel_sum (rel_prod IA (=)) (rel_prod IU (=)))) ===> + rel_prod (E ===> (=)) (rel_prod (IA ===> (=)) (IU ===> (=)))) + trace_core' trace_core'" + unfolding trace_core'_def by transfer_prover + +definition trace_core_eq + :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core + \ ('s_core', 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core + \ 'event set \ 'iadv_core set \ 'iusr_core set + \ 's_core spmf \ 's_core' spmf \ bool" where + "trace_core_eq core1 core2 E IA IU p q \ + (\tr. set tr \ E <+> (IA \ UNIV) <+> (IU \ UNIV) \ + rel_prod (eq_onp (\e. e \ E) ===> (=)) (rel_prod (eq_onp (\ia. ia \ IA) ===> (=)) (eq_onp (\iu. iu \ IU) ===> (=))) + (trace_core' core1 p tr) (trace_core' core2 q tr))" + +end + +lemma trace_core_eqD: + assumes "trace_core_eq core1 core2 E IA IU p q" + and "set tr \ E <+> (IA \ UNIV) <+> (IU \ UNIV)" + shows trace_core_eqD_cpoke: + "e \ E \ fst (trace_core' core1 p tr) e = fst (trace_core' core2 q tr) e" + and trace_core_eqD_cfunc_adv: + "ia \ IA \ fst (snd (trace_core' core1 p tr)) ia = fst (snd (trace_core' core2 q tr)) ia" + and trace_core_eqD_cfunc_usr: + "iu \ IU \ snd (snd (trace_core' core1 p tr)) iu = snd (snd (trace_core' core2 q tr)) iu" + using assms by(auto simp add: trace_core_eq_def rel_fun_def eq_onp_def rel_prod_sel) + +lemma trace_core_eqI: + assumes "\tr e. \ set tr \ E <+> (IA \ UNIV) <+> (IU \ UNIV); e \ E \ + \ fst (trace_core' core1 p tr) e = fst (trace_core' core2 q tr) e" + and "\tr ia. \ set tr \ E <+> (IA \ UNIV) <+> (IU \ UNIV); ia \ IA \ + \ fst (snd (trace_core' core1 p tr)) ia = fst (snd (trace_core' core2 q tr)) ia" + and "\tr iu. \ set tr \ E <+> (IA \ UNIV) <+> (IU \ UNIV); iu \ IU \ + \ snd (snd (trace_core' core1 p tr)) iu = snd (snd (trace_core' core2 q tr)) iu" + shows "trace_core_eq core1 core2 E IA IU p q" + using assms by(auto simp add: trace_core_eq_def rel_fun_def eq_onp_def rel_prod_sel) + +lemma trace_core_return_pmf_None [simp]: + "trace_core' core (return_pmf None) tr = (\_. 0, \_. return_pmf None, \_. return_pmf None)" + by(induction tr)(simp_all add: trace_core'.simps split: sum.split) + +lemma rel_core'_into_trace_core_eq: "trace_core_eq core core' E IA IU p q" + if "rel_core' S (eq_onp (\e. e \ E)) (eq_onp (\ia. ia \ IA)) (eq_onp (\iu. iu \ IU)) (=) (=) core core'" + "rel_spmf S p q" + using trace_core'_parametric[THEN rel_funD, THEN rel_funD, OF that] + unfolding trace_core_eq_def + apply(intro strip) + subgoal for tr + apply(simp add: eq_onp_True[symmetric] prod.rel_eq_onp sum.rel_eq_onp list.rel_eq_onp) + apply(auto 4 3 simp add: eq_onp_def list_all_iff dest: rel_funD[where x=tr and y=tr]) + done + done + +lemma trace_core_eq_simI: + fixes core1 :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and core2 :: "('s_core', 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and S :: "'s_core spmf \ 's_core' spmf \ bool" + assumes start: "S p q" + and step_cpoke: "\p q e. \ S p q; e \ E \ \ + weight_spmf (bind_spmf p (\s. cpoke core1 s e)) = weight_spmf (bind_spmf q (\s. cpoke core2 s e))" + and sim_cpoke: "\p q e. \ S p q; e \ E \ \ + S (mk_lossless (bind_spmf p (\s. cpoke core1 s e))) (mk_lossless (bind_spmf q (\s. cpoke core2 s e)))" + and step_cfunc_adv: "\p q ia. \ S p q; ia \ IA \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_adv core1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (cfunc_adv core2 s2 ia))" + and sim_cfunc_adv: "\p q ia s1 s2 s1' s2' oa. \ S p q; ia \ IA; + s1 \ set_spmf p; s2 \ set_spmf q; (oa, s1') \ set_spmf (cfunc_adv core1 s1 ia); (oa, s2') \ set_spmf (cfunc_adv core2 s2 ia) \ + \ S (cond_spmf_fst (bind_spmf p (\s1. cfunc_adv core1 s1 ia)) oa) (cond_spmf_fst (bind_spmf q (\s2. cfunc_adv core2 s2 ia)) oa)" + and step_cfunc_usr: "\p q iu. \ S p q; iu \ IU \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_usr core1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (cfunc_usr core2 s2 iu))" + and sim_cfunc_usr: "\p q iu s1 s2 s1' s2' ou. \ S p q; iu \ IU; + s1 \ set_spmf p; s2 \ set_spmf q; (ou, s1') \ set_spmf (cfunc_usr core1 s1 iu); (ou, s2') \ set_spmf (cfunc_usr core2 s2 iu) \ + \ S (cond_spmf_fst (bind_spmf p (\s1. cfunc_usr core1 s1 iu)) ou) (cond_spmf_fst (bind_spmf q (\s2. cfunc_usr core2 s2 iu)) ou)" + shows "trace_core_eq core1 core2 E IA IU p q" +proof(rule trace_core_eqI) + fix tr :: " ('event + 'iadv_core \ 'oadv_core + 'iusr_core \ 'ousr_core) list" + assume "set tr \ E <+> IA \ UNIV <+> IU \ UNIV" + then have "(\e\E. fst (trace_core' core1 p tr) e = fst (trace_core' core2 q tr) e) \ + (\ia\IA. fst (snd (trace_core' core1 p tr)) ia = fst (snd (trace_core' core2 q tr)) ia) \ + (\iu\IU. snd (snd (trace_core' core1 p tr)) iu = snd (snd (trace_core' core2 q tr)) iu)" + using start + proof(induction tr arbitrary: p q) + case Nil + then show ?case by(simp add: step_cpoke step_cfunc_adv step_cfunc_usr) + next + case (Cons a tr) + from Cons.prems(1) have tr: "set tr \ E <+> IA \ UNIV <+> IU \ UNIV" by simp + from Cons.prems(1) + consider (cpoke) e where "a = Inl e" "e \ E" + | (cfunc_adv) ia oa where "a = Inr (Inl (ia, oa))" "ia \ IA" + | (cfunc_usr) iu ou where "a = Inr (Inr (iu, ou))" "iu \ IU" by auto + then show ?case + proof cases + case cpoke + then show ?thesis using tr Cons.prems(2) by(auto simp add: sim_cpoke intro!: Cons.IH) + next + case cfunc_adv + let ?p = "bind_spmf p (\s1. cfunc_adv core1 s1 ia)" + let ?q = "bind_spmf q (\s2. cfunc_adv core2 s2 ia)" + show ?thesis + proof(cases "oa \ fst ` set_spmf ?p") + case True + with step_cfunc_adv[OF Cons.prems(2) cfunc_adv(2), THEN arg_cong[where f=set_spmf]] + have "oa \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) + then show ?thesis using True Cons.prems cfunc_adv + by(clarsimp)(rule Cons.IH; blast intro: sim_cfunc_adv) + next + case False + hence "cond_spmf_fst ?p oa = return_pmf None" by simp + moreover + from step_cfunc_adv[OF Cons.prems(2) cfunc_adv(2), THEN arg_cong[where f=set_spmf]] False + have oa': "oa \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) simp + hence "cond_spmf_fst ?q oa = return_pmf None" by simp + ultimately show ?thesis using cfunc_adv by(simp del: cond_spmf_fst_eq_return_None) + qed + next + case cfunc_usr + let ?p = "bind_spmf p (\s1. cfunc_usr core1 s1 iu)" + let ?q = "bind_spmf q (\s2. cfunc_usr core2 s2 iu)" + show ?thesis + proof(cases "ou \ fst ` set_spmf ?p") + case True + with step_cfunc_usr[OF Cons.prems(2) cfunc_usr(2), THEN arg_cong[where f=set_spmf]] + have "ou \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) + then show ?thesis using True Cons.prems cfunc_usr + by(clarsimp)(rule Cons.IH; blast intro: sim_cfunc_usr) + next + case False + hence "cond_spmf_fst ?p ou = return_pmf None" by simp + moreover + from step_cfunc_usr[OF Cons.prems(2) cfunc_usr(2), THEN arg_cong[where f=set_spmf]] False + have oa': "ou \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) simp + hence "cond_spmf_fst ?q ou = return_pmf None" by simp + ultimately show ?thesis using cfunc_usr by(simp del: cond_spmf_fst_eq_return_None) + qed + qed + qed + then show "e \ E \ fst (trace_core' core1 p tr) e = fst (trace_core' core2 q tr) e" + and "ia \ IA \ fst (snd (trace_core' core1 p tr)) ia = fst (snd (trace_core' core2 q tr)) ia" + and "iu \ IU \ snd (snd (trace_core' core1 p tr)) iu = snd (snd (trace_core' core2 q tr)) iu" + for e ia iu by blast+ +qed + +context + fixes core :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" +begin + +fun trace_core_aux + :: "'s_core spmf \ ('event + 'iadv_core \ 'oadv_core + 'iusr_core \ 'ousr_core) list \ 's_core spmf" where + "trace_core_aux p [] = p" +| "trace_core_aux p (Inl e # tr) = trace_core_aux (mk_lossless (bind_spmf p (\s. cpoke core s e))) tr" +| "trace_core_aux p (Inr (Inl (ia, oa)) # tr) = trace_core_aux (cond_spmf_fst (bind_spmf p (\s. cfunc_adv core s ia)) oa) tr" +| "trace_core_aux p (Inr (Inr (iu, ou)) # tr) = trace_core_aux (cond_spmf_fst (bind_spmf p (\s. cfunc_usr core s iu)) ou) tr" + +end + +lemma trace_core_conv_trace_core_aux: + "trace_core' core p tr = + (\e. weight_spmf (bind_spmf (trace_core_aux core p tr) (\s. cpoke core s e)), + \ia. bind_spmf (trace_core_aux core p tr) (\s. map_spmf fst (cfunc_adv core s ia)), + \iu. bind_spmf (trace_core_aux core p tr) (\s. map_spmf fst (cfunc_usr core s iu)))" + by(induction p tr rule: trace_core_aux.induct) simp_all + +lemma trace_core_aux_append: + "trace_core_aux core p (tr @ tr') = trace_core_aux core (trace_core_aux core p tr) tr'" + by(induction p tr rule: trace_core_aux.induct) auto + +inductive trace_core_closure + :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core + \ ('s_core', 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core + \ 'event set \ 'iadv_core set \ 'iusr_core set + \ 's_core spmf \ 's_core' spmf \ 's_core spmf \ 's_core' spmf \ bool" + for core1 core2 E IA IU p q where + "trace_core_closure core1 core2 E IA IU p q (trace_core_aux core1 p tr) (trace_core_aux core2 q tr)" + if "set tr \ E <+> IA \ UNIV <+> IU \ UNIV" + +lemma trace_core_closure_start: "trace_core_closure core1 core2 E IA IU p q p q" + by(simp add: trace_core_closure.simps exI[where x="[]"]) + +lemma trace_core_closure_step: + assumes "trace_core_eq core1 core2 E IA IU p q" + and "trace_core_closure core1 core2 E IA IU p q p' q'" + shows trace_core_closure_step_cpoke: + "e \ E \ weight_spmf (bind_spmf p' (\s. cpoke core1 s e)) = weight_spmf (bind_spmf q' (\s. cpoke core2 s e))" + (is "PROP ?thesis1") + and trace_core_closure_step_cfunc_adv: + "ia \ IA \ bind_spmf p' (\s1. map_spmf fst (cfunc_adv core1 s1 ia)) = bind_spmf q' (\s2. map_spmf fst (cfunc_adv core2 s2 ia))" + (is "PROP ?thesis2") + and trace_core_closure_step_cfunc_usr: + "iu \ IU \ bind_spmf p' (\s1. map_spmf fst (cfunc_usr core1 s1 iu)) = bind_spmf q' (\s2. map_spmf fst (cfunc_usr core2 s2 iu))" + (is "PROP ?thesis3") +proof - + from assms(2) obtain tr where p: "p' = trace_core_aux core1 p tr" + and q: "q' = trace_core_aux core2 q tr" + and tr: "set tr \ E <+> IA \ UNIV <+> IU \ UNIV" by cases + from trace_core_eqD[OF assms(1) tr] p q + show "PROP ?thesis1" and "PROP ?thesis2" "PROP ?thesis3" + by(simp_all add: trace_core_conv_trace_core_aux) +qed + +lemma trace_core_closure_sim: + fixes core1 core2 E IA IU p q + defines "S \ trace_core_closure core1 core2 E IA IU p q" + assumes "S p' q'" + shows trace_core_closure_sim_cpoke: + "e \ E \ S (mk_lossless (bind_spmf p' (\s. cpoke core1 s e))) (mk_lossless (bind_spmf q' (\s. cpoke core2 s e)))" + (is "PROP ?thesis1") + and trace_core_closure_sim_cfunc_adv: "ia \ IA + \ S (cond_spmf_fst (bind_spmf p' (\s1. cfunc_adv core1 s1 ia)) oa) + (cond_spmf_fst (bind_spmf q' (\s2. cfunc_adv core2 s2 ia)) oa)" + (is "PROP ?thesis2") + and trace_core_closure_sim_cfunc_usr: "iu \ IU + \ S (cond_spmf_fst (bind_spmf p' (\s1. cfunc_usr core1 s1 iu)) ou) + (cond_spmf_fst (bind_spmf q' (\s2. cfunc_usr core2 s2 iu)) ou)" + (is "PROP ?thesis3") +proof - + from assms(2) obtain tr where p: "p' = trace_core_aux core1 p tr" + and q: "q' = trace_core_aux core2 q tr" + and tr: "set tr \ E <+> IA \ UNIV <+> IU \ UNIV" unfolding S_def by cases + show "PROP ?thesis1" using p q tr + by(auto simp add: S_def trace_core_closure.simps trace_core_aux_append intro!: exI[where x="tr @ [Inl _]"]) + show "PROP ?thesis2" using p q tr + by(auto simp add: S_def trace_core_closure.simps trace_core_aux_append intro!: exI[where x="tr @ [Inr (Inl (_, _))]"]) + show "PROP ?thesis3" using p q tr + by(auto simp add: S_def trace_core_closure.simps trace_core_aux_append intro!: exI[where x="tr @ [Inr (Inr (_, _))]"]) +qed + +proposition trace_core_eq_complete: + assumes "trace_core_eq core1 core2 E IA IU p q" + obtains S + where "S p q" + and "\p q e. \ S p q; e \ E \ \ + weight_spmf (bind_spmf p (\s. cpoke core1 s e)) = weight_spmf (bind_spmf q (\s. cpoke core2 s e))" + and "\p q e. \ S p q; e \ E \ \ + S (mk_lossless (bind_spmf p (\s. cpoke core1 s e))) (mk_lossless (bind_spmf q (\s. cpoke core2 s e)))" + and "\p q ia. \ S p q; ia \ IA \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_adv core1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (cfunc_adv core2 s2 ia))" + and "\p q ia oa. \ S p q; ia \ IA \ + \ S (cond_spmf_fst (bind_spmf p (\s1. cfunc_adv core1 s1 ia)) oa) (cond_spmf_fst (bind_spmf q (\s2. cfunc_adv core2 s2 ia)) oa)" + and "\p q iu. \ S p q; iu \ IU \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_usr core1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (cfunc_usr core2 s2 iu))" + and "\p q iu ou. \ S p q; iu \ IU \ + \ S (cond_spmf_fst (bind_spmf p (\s1. cfunc_usr core1 s1 iu)) ou) (cond_spmf_fst (bind_spmf q (\s2. cfunc_usr core2 s2 iu)) ou)" +proof - + show thesis + by(rule that[where S="trace_core_closure core1 core2 E IA IU p q"]) + (auto intro: trace_core_closure_start trace_core_closure_step[OF assms] trace_core_closure_sim (* trace_core_closure_weight[OF assms] *)) +qed + + + +type_synonym ('event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) trace_rest = + "('iadv_rest \ 'oadv_rest \ 'event list + 'iusr_rest \ 'ousr_rest \ 'event list) list + \ ('iadv_rest \ ('oadv_rest \ 'event list) spmf) + \ ('iusr_rest \ ('ousr_rest \ 'event list) spmf)" + +context + fixes rest :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" +begin + +primrec trace_rest' :: "'s_rest spmf \ ('event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest) trace_rest" where + "trace_rest' S [] = + (\ia. bind_spmf S (\s. map_spmf fst (rfunc_adv rest s ia)), + \iu. bind_spmf S (\s. map_spmf fst (rfunc_usr rest s iu)))" +| "trace_rest' S (obs # tr) = (case obs of + Inl (ia, oa) \ trace_rest' (cond_spmf_fst (bind_spmf S (\s. rfunc_adv rest s ia)) oa) tr + | Inr (iu, ou) \ trace_rest' (cond_spmf_fst (bind_spmf S (\s. rfunc_usr rest s iu)) ou) tr)" + +end + +declare trace_rest'.simps [simp del] +case_of_simps trace_rest'_unfold: trace_rest'.simps +simps_of_case trace_rest'_simps [simp]: trace_rest'_unfold + +context includes lifting_syntax begin + +lemma trace_rest'_parametric [transfer_rule]: + "(rel_rest' S (=) IA IU (=) (=) M ===> rel_spmf S ===> + list_all2 (rel_sum (rel_prod IA (=)) (rel_prod IU (=))) ===> + rel_prod (IA ===> (=)) (IU ===> (=))) + trace_rest' trace_rest'" + unfolding trace_rest'_def by transfer_prover + +definition trace_rest_eq + :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more1) rest_scheme + \ ('s_rest', 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more2) rest_scheme + \ 'iadv_rest set \ 'iusr_rest set + \ 's_rest spmf \ 's_rest' spmf \ bool" where + "trace_rest_eq rest1 rest2 IA IU p q \ + (\tr. set tr \ (IA \ UNIV) <+> (IU \ UNIV) \ + rel_prod (eq_onp (\ia. ia \ IA) ===> (=)) (eq_onp (\iu. iu \ IU) ===> (=)) + (trace_rest' rest1 p tr) (trace_rest' rest2 q tr))" + +end + +lemma trace_rest_eqD: + assumes "trace_rest_eq rest1 rest2 IA IU p q" + and "set tr \ (IA \ UNIV) <+> (IU \ UNIV)" + shows trace_rest_eqD_rfunc_adv: + "ia \ IA \ fst (trace_rest' rest1 p tr) ia = fst (trace_rest' rest2 q tr) ia" + and trace_rest_eqD_rfunc_usr: + "iu \ IU \ snd (trace_rest' rest1 p tr) iu = snd (trace_rest' rest2 q tr) iu" + using assms by(auto simp add: trace_rest_eq_def rel_fun_def rel_prod_sel eq_onp_def) + +lemma trace_rest_eqI: + assumes "\tr ia. \ set tr \ (IA \ UNIV) <+> (IU \ UNIV); ia \ IA \ + \ fst (trace_rest' rest1 p tr) ia = fst (trace_rest' rest2 q tr) ia" + and "\tr iu. \ set tr \ (IA \ UNIV) <+> (IU \ UNIV); iu \ IU \ + \ snd (trace_rest' rest1 p tr) iu = snd (trace_rest' rest2 q tr) iu" + shows "trace_rest_eq rest1 rest2 IA IU p q" + using assms by(auto simp add: trace_rest_eq_def rel_fun_def eq_onp_def rel_prod_sel) + +lemma trace_rest_return_pmf_None [simp]: + "trace_rest' rest (return_pmf None) tr = (\_. return_pmf None, \_. return_pmf None)" + by(induction tr)(simp_all add: trace_rest'.simps split: sum.split) + +lemma rel_rest'_into_trace_rest_eq: "trace_rest_eq rest rest' IA IU p q" + if "rel_rest' S (=) (eq_onp (\ia. ia \ IA)) (eq_onp (\iu. iu \ IU)) (=) (=) M rest rest'" + "rel_spmf S p q" + using trace_rest'_parametric[THEN rel_funD, THEN rel_funD, OF that] + unfolding trace_rest_eq_def + apply(intro strip) + subgoal for tr + apply(simp add: eq_onp_True[symmetric] prod.rel_eq_onp sum.rel_eq_onp list.rel_eq_onp) + apply(auto 4 3 simp add: eq_onp_def list_all_iff dest: rel_funD[where x=tr and y=tr]) + done + done + +lemma trace_rest_eq_simI: + fixes rest1 :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" + and rest2 :: "('s_rest', 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" + and S :: "'s_rest spmf \ 's_rest' spmf \ bool" + assumes start: "S p q" + and step_rfunc_adv: "\p q ia. \ S p q; ia \ IA \ \ + bind_spmf p (\s1. map_spmf fst (rfunc_adv rest1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (rfunc_adv rest2 s2 ia))" + and sim_rfunc_adv: "\p q ia s1 s2 s1' s2' oa. \ S p q; ia \ IA; + s1 \ set_spmf p; s2 \ set_spmf q; (oa, s1') \ set_spmf (rfunc_adv rest1 s1 ia); (oa, s2') \ set_spmf (rfunc_adv rest2 s2 ia) \ + \ S (cond_spmf_fst (bind_spmf p (\s1. rfunc_adv rest1 s1 ia)) oa) (cond_spmf_fst (bind_spmf q (\s2. rfunc_adv rest2 s2 ia)) oa)" + and step_rfunc_usr: "\p q iu. \ S p q; iu \ IU \ \ + bind_spmf p (\s1. map_spmf fst (rfunc_usr rest1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (rfunc_usr rest2 s2 iu))" + and sim_rfunc_usr: "\p q iu s1 s2 s1' s2' ou. \ S p q; iu \ IU; + s1 \ set_spmf p; s2 \ set_spmf q; (ou, s1') \ set_spmf (rfunc_usr rest1 s1 iu); (ou, s2') \ set_spmf (rfunc_usr rest2 s2 iu) \ + \ S (cond_spmf_fst (bind_spmf p (\s1. rfunc_usr rest1 s1 iu)) ou) (cond_spmf_fst (bind_spmf q (\s2. rfunc_usr rest2 s2 iu)) ou)" + shows "trace_rest_eq rest1 rest2 IA IU p q" +proof(rule trace_rest_eqI) + fix tr :: "('iadv_rest \ 'oadv_rest \ 'event list + 'iusr_rest \ 'ousr_rest \ 'event list) list" + assume "set tr \ IA \ UNIV <+> IU \ UNIV" + then have "(\ia\IA. fst (trace_rest' rest1 p tr) ia = fst (trace_rest' rest2 q tr) ia) \ + (\iu\IU. snd (trace_rest' rest1 p tr) iu = snd (trace_rest' rest2 q tr) iu)" + using start + proof(induction tr arbitrary: p q) + case Nil + then show ?case by(simp add: step_rfunc_adv step_rfunc_usr) + next + case (Cons a tr) + from Cons.prems(1) have tr: "set tr \ IA \ UNIV <+> IU \ UNIV" by simp + from Cons.prems(1) + consider (rfunc_adv) ia oa where "a = Inl (ia, oa)" "ia \ IA" + | (rfunc_usr) iu ou where "a = Inr (iu, ou)" "iu \ IU" by auto + then show ?case + proof cases + case rfunc_adv + let ?p = "bind_spmf p (\s1. rfunc_adv rest1 s1 ia)" + let ?q = "bind_spmf q (\s2. rfunc_adv rest2 s2 ia)" + show ?thesis + proof(cases "oa \ fst ` set_spmf ?p") + case True + with step_rfunc_adv[OF Cons.prems(2) rfunc_adv(2), THEN arg_cong[where f=set_spmf]] + have "oa \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) + then show ?thesis using True Cons.prems rfunc_adv + by(clarsimp)(rule Cons.IH; blast intro: sim_rfunc_adv) + next + case False + hence "cond_spmf_fst ?p oa = return_pmf None" by simp + moreover + from step_rfunc_adv[OF Cons.prems(2) rfunc_adv(2), THEN arg_cong[where f=set_spmf]] False + have oa': "oa \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) simp + hence "cond_spmf_fst ?q oa = return_pmf None" by simp + ultimately show ?thesis using rfunc_adv by(simp del: cond_spmf_fst_eq_return_None) + qed + next + case rfunc_usr + let ?p = "bind_spmf p (\s1. rfunc_usr rest1 s1 iu)" + let ?q = "bind_spmf q (\s2. rfunc_usr rest2 s2 iu)" + show ?thesis + proof(cases "ou \ fst ` set_spmf ?p") + case True + with step_rfunc_usr[OF Cons.prems(2) rfunc_usr(2), THEN arg_cong[where f=set_spmf]] + have "ou \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) + then show ?thesis using True Cons.prems rfunc_usr + by(clarsimp)(rule Cons.IH; blast intro: sim_rfunc_usr) + next + case False + hence "cond_spmf_fst ?p ou = return_pmf None" by simp + moreover + from step_rfunc_usr[OF Cons.prems(2) rfunc_usr(2), THEN arg_cong[where f=set_spmf]] False + have oa': "ou \ fst ` set_spmf ?q" + unfolding set_map_spmf[symmetric] by(simp only: map_bind_spmf o_def) simp + hence "cond_spmf_fst ?q ou = return_pmf None" by simp + ultimately show ?thesis using rfunc_usr by(simp del: cond_spmf_fst_eq_return_None) + qed + qed + qed + then show "ia \ IA \ fst (trace_rest' rest1 p tr) ia = fst (trace_rest' rest2 q tr) ia" + and "iu \ IU \ snd (trace_rest' rest1 p tr) iu = snd (trace_rest' rest2 q tr) iu" + for ia iu by blast+ +qed + +context + fixes rest :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" +begin + +fun trace_rest_aux + :: "'s_rest spmf \ ('iadv_rest \ 'oadv_rest \ 'event list + 'iusr_rest \ 'ousr_rest \ 'event list) list \ 's_rest spmf" where + "trace_rest_aux p [] = p" +| "trace_rest_aux p (Inl (ia, oaes) # tr) = trace_rest_aux (cond_spmf_fst (bind_spmf p (\s. rfunc_adv rest s ia)) oaes) tr" +| "trace_rest_aux p (Inr (iu, oues) # tr) = trace_rest_aux (cond_spmf_fst (bind_spmf p (\s. rfunc_usr rest s iu)) oues) tr" + +end + +lemma trace_rest_conv_trace_rest_aux: + "trace_rest' rest p tr = + (\ia. bind_spmf (trace_rest_aux rest p tr) (\s. map_spmf fst (rfunc_adv rest s ia)), + \iu. bind_spmf (trace_rest_aux rest p tr) (\s. map_spmf fst (rfunc_usr rest s iu)))" + by(induction p tr rule: trace_rest_aux.induct) simp_all + +lemma trace_rest_aux_append: + "trace_rest_aux rest p (tr @ tr') = trace_rest_aux rest (trace_rest_aux rest p tr) tr'" + by(induction p tr rule: trace_rest_aux.induct) auto + +inductive trace_rest_closure + :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme + \ ('s_rest', 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more') rest_scheme + \ 'iadv_rest set \ 'iusr_rest set + \ 's_rest spmf \ 's_rest' spmf \ 's_rest spmf \ 's_rest' spmf \ bool" + for rest1 rest2 IA IU p q where + "trace_rest_closure rest1 rest2 IA IU p q (trace_rest_aux rest1 p tr) (trace_rest_aux rest2 q tr)" + if "set tr \ IA \ UNIV <+> IU \ UNIV" + +lemma trace_rest_closure_start: "trace_rest_closure rest1 rest2 IA IU p q p q" + by(simp add: trace_rest_closure.simps exI[where x="[]"]) + +lemma trace_rest_closure_step: + assumes "trace_rest_eq rest1 rest2 IA IU p q" + and "trace_rest_closure rest1 rest2 IA IU p q p' q'" + shows trace_rest_closure_step_rfunc_adv: + "ia \ IA \ bind_spmf p' (\s1. map_spmf fst (rfunc_adv rest1 s1 ia)) = bind_spmf q' (\s2. map_spmf fst (rfunc_adv rest2 s2 ia))" + (is "PROP ?thesis1") + and trace_rest_closure_step_rfunc_usr: + "iu \ IU \ bind_spmf p' (\s1. map_spmf fst (rfunc_usr rest1 s1 iu)) = bind_spmf q' (\s2. map_spmf fst (rfunc_usr rest2 s2 iu))" + (is "PROP ?thesis2") +proof - + from assms(2) obtain tr where p: "p' = trace_rest_aux rest1 p tr" + and q: "q' = trace_rest_aux rest2 q tr" + and tr: "set tr \ IA \ UNIV <+> IU \ UNIV" by cases + from trace_rest_eqD[OF assms(1) tr] p q + show "PROP ?thesis1" and "PROP ?thesis2" + by(simp_all add: trace_rest_conv_trace_rest_aux) +qed + +lemma trace_rest_closure_sim: + fixes rest1 rest2 IA IU p q + defines "S \ trace_rest_closure rest1 rest2 IA IU p q" + assumes "S p' q'" + shows trace_rest_closure_sim_rfunc_adv: "ia \ IA + \ S (cond_spmf_fst (bind_spmf p' (\s1. rfunc_adv rest1 s1 ia)) oa) + (cond_spmf_fst (bind_spmf q' (\s2. rfunc_adv rest2 s2 ia)) oa)" + (is "PROP ?thesis1") + and trace_rest_closure_sim_rfunc_usr: "iu \ IU + \ S (cond_spmf_fst (bind_spmf p' (\s1. rfunc_usr rest1 s1 iu)) ou) + (cond_spmf_fst (bind_spmf q' (\s2. rfunc_usr rest2 s2 iu)) ou)" + (is "PROP ?thesis2") +proof - + from assms(2) obtain tr where p: "p' = trace_rest_aux rest1 p tr" + and q: "q' = trace_rest_aux rest2 q tr" + and tr: "set tr \ IA \ UNIV <+> IU \ UNIV" unfolding S_def by cases + show "PROP ?thesis1" using p q tr + by(auto simp add: S_def trace_rest_closure.simps trace_rest_aux_append intro!: exI[where x="tr @ [Inl (_, _)]"]) + show "PROP ?thesis2" using p q tr + by(auto simp add: S_def trace_rest_closure.simps trace_rest_aux_append intro!: exI[where x="tr @ [Inr (_, _)]"]) +qed + +proposition trace_rest_eq_complete: + assumes "trace_rest_eq rest1 rest2 IA IU p q" + obtains S + where "S p q" + and "\p q ia. \ S p q; ia \ IA \ \ + bind_spmf p (\s1. map_spmf fst (rfunc_adv rest1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (rfunc_adv rest2 s2 ia))" + and "\p q ia oa. \ S p q; ia \ IA \ + \ S (cond_spmf_fst (bind_spmf p (\s1. rfunc_adv rest1 s1 ia)) oa) (cond_spmf_fst (bind_spmf q (\s2. rfunc_adv rest2 s2 ia)) oa)" + and "\p q iu. \ S p q; iu \ IU \ \ + bind_spmf p (\s1. map_spmf fst (rfunc_usr rest1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (rfunc_usr rest2 s2 iu))" + and "\p q iu ou. \ S p q; iu \ IU \ + \ S (cond_spmf_fst (bind_spmf p (\s1. rfunc_usr rest1 s1 iu)) ou) (cond_spmf_fst (bind_spmf q (\s2. rfunc_usr rest2 s2 iu)) ou)" +proof - + show thesis + by(rule that[where S="trace_rest_closure rest1 rest2 IA IU p q"]) + (auto intro: trace_rest_closure_start trace_rest_closure_step[OF assms] trace_rest_closure_sim (* trace_rest_closure_weight[OF assms] *)) +qed + +definition callee_of_core + :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core + \ ('s_core, 'event + 'iadv_core + 'iusr_core, unit + 'oadv_core + 'ousr_core) oracle'" where + "callee_of_core core = + map_fun id (map_fun id (map_spmf (Pair ()))) (cpoke core) \\<^sub>O cfunc_adv core \\<^sub>O cfunc_usr core" + +lemma callee_of_core_simps [simp]: + "callee_of_core core s (Inl e) = map_spmf (Pair (Inl ())) (cpoke core s e)" + "callee_of_core core s (Inr (Inl iadv_core)) = map_spmf (apfst (Inr \ Inl)) (cfunc_adv core s iadv_core)" + "callee_of_core core s (Inr (Inr iusr_core)) = map_spmf (apfst (Inr \ Inr)) (cfunc_usr core s iusr_core)" + by(simp_all add: callee_of_core_def spmf.map_comp o_def apfst_def prod.map_comp id_def) + +lemma WT_callee_of_core [WT_intro]: + assumes WT: "WT_core \_adv \_usr I core" + and I: "I s" + shows "\_full \\<^sub>\ (\_adv \\<^sub>\ \_usr) \c callee_of_core core s \" + apply(rule WT_calleeI) + subgoal for x y s' using I WT_coreD[OF WT] + by(auto simp add: callee_of_core_def plus_oracle_def split!: sum.splits) + done + +lemma WT_core_callee_invariant_on [WT_intro]: + assumes WT: "WT_core \_adv \_usr I core" + shows "callee_invariant_on (callee_of_core core) I (\_full \\<^sub>\ (\_adv \\<^sub>\ \_usr))" + apply unfold_locales + subgoal for s x y s' by(auto simp add: callee_of_core_def plus_oracle_def split!: sum.splits dest: WT_coreD[OF assms]) + subgoal by(rule WT_callee_of_core[OF WT]) + done + +definition callee_of_rest + :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme + \ ('s_rest, 'iadv_rest + 'iusr_rest, 'oadv_rest \ 'event list + 'ousr_rest \ 'event list) oracle'" where + "callee_of_rest rest = rfunc_adv rest \\<^sub>O rfunc_usr rest" + +lemma callee_of_rest_simps [simp]: + "callee_of_rest rest s (Inl iadv_rest) = map_spmf (apfst Inl) (rfunc_adv rest s iadv_rest)" + "callee_of_rest rest s (Inr iusr_rest) = map_spmf (apfst Inr) (rfunc_usr rest s iusr_rest)" + by(simp_all add: callee_of_rest_def) + +lemma WT_callee_of_rest [WT_intro]: + assumes WT: "WT_rest \_adv \_usr I rest" + and I: "I s" + shows "e\ \_adv \\<^sub>\ e\ \_usr \c callee_of_rest rest s \" + apply(rule WT_calleeI) + subgoal for x y s' using I WT_restD[OF WT] + by(auto simp add: callee_of_core_def plus_oracle_def split!: sum.splits) + done + +fun fuse_callee + :: "('iadv_core + 'iadv_rest) + ('iusr_core + 'iusr_rest) \ + (('oadv_core + 'oadv_rest) + ('ousr_core + 'ousr_rest), + ('event + 'iadv_core + 'iusr_core) + ('iadv_rest + 'iusr_rest), + (unit + 'oadv_core + 'ousr_core) + ('oadv_rest \ 'event list + 'ousr_rest \ 'event list)) gpv" + where + "fuse_callee (Inl (Inl iadv_core)) = Pause (Inl (Inr (Inl iadv_core))) (\x. case x of + Inl (Inr (Inl oadv_core)) \ Done (Inl (Inl oadv_core)) + | _ \ Fail)" +| "fuse_callee (Inl (Inr iadv_rest)) = Pause (Inr (Inl iadv_rest)) (\x. case x of + Inr (Inl (oadv_rest, es)) \ bind_gpv (pauses (map (Inl \ Inl) es)) (\_. Done (Inl (Inr oadv_rest))) + | _ \ Fail)" +| "fuse_callee (Inr (Inl iusr_core)) = Pause (Inl (Inr (Inr iusr_core))) (\x. case x of + Inl (Inr (Inr oadv_core)) \ Done (Inr (Inl oadv_core)))" +| "fuse_callee (Inr (Inr iusr_rest)) = Pause (Inr (Inr iusr_rest)) (\x. case x of + Inr (Inr (ousr_rest, es)) \ bind_gpv (pauses (map (Inl \ Inl) es)) (\_. Done (Inr (Inr ousr_rest))))" + +case_of_simps fuse_callee_case: fuse_callee.simps + +(* parametric_constant fuse_callee_parametric [transfer_rule]: fuse_callee_case *) + +definition fuse_converter + :: "(('iadv_core + 'iadv_rest) + ('iusr_core + 'iusr_rest), + ('oadv_core + 'oadv_rest) + ('ousr_core + 'ousr_rest), + ('event + 'iadv_core + 'iusr_core) + ('iadv_rest + 'iusr_rest), + (unit + 'oadv_core + 'ousr_core) + ('oadv_rest \ 'event list + 'ousr_rest \ 'event list)) converter" + where + "fuse_converter = converter_of_callee (stateless_callee fuse_callee) ()" + +lemma fuse_converter: + "resource_of_oracle (fused_resource.fuse core rest) (s_core, s_rest) = + fuse_converter \ (resource_of_oracle (callee_of_core core) s_core \ resource_of_oracle (callee_of_rest rest) s_rest)" + unfolding fuse_converter_def resource_of_parallel_oracle[symmetric] attach_CNV_RES attach_stateless_callee resource_of_oracle_extend_state_oracle +proof(rule arg_cong2[where f=resource_of_oracle]; clarsimp simp add: fun_eq_iff) + interpret fused_resource core core_init for core_init . + have "foldl_spmf (map_fun id (map_fun (Inl \ Inl) id) (map_fun id (map_fun id (map_spmf snd)) (callee_of_core core \\<^sub>O callee_of_rest rest))) (return_spmf (s_core, s_rest)) xs + = map_spmf (\s_core. (s_core, s_rest)) (foldl_spmf (cpoke core) (return_spmf s_core) xs)" for s_core s_rest xs + by(induction xs arbitrary: s_core) + (simp_all add: spmf.map_comp foldl_spmf_Cons' map_bind_spmf bind_map_spmf o_def del: foldl_spmf_Cons) + then show "fuse rest (s_core, s_rest) q = exec_gpv (callee_of_core core \\<^sub>O callee_of_rest rest) (fuse_callee q) (s_core, s_rest)" + for s_core s_rest q + by(cases q rule: fuse_callee.cases; clarsimp simp add: map_bind_spmf bind_map_spmf exec_gpv_bind exec_gpv_pauses intro!: bind_spmf_cong[OF refl]; simp add: map_spmf_conv_bind_spmf[symmetric]) +qed + +lemma trace_eq_callee_of_coreI: + "trace_callee_eq (callee_of_core core1) (callee_of_core core2) (E <+> IA <+> IU) p q" + if "trace_core_eq core1 core2 E IA IU p q" +proof - + from that obtain S_core + where core_start: "S_core p q" + and step_cpoke: "\p q e. S_core p q \ e \ E + \ weight_spmf (bind_spmf p (\s. cpoke core1 s e)) = weight_spmf (bind_spmf q (\s. cpoke core2 s e))" + and sim_cpoke: "\p q e. S_core p q \ e \ E + \ S_core (mk_lossless (bind_spmf p (\s. cpoke core1 s e))) (mk_lossless (bind_spmf q (\s. cpoke core2 s e)))" + and step_cfunc_adv: "\p q ia. \ S_core p q; ia \ IA \ + \ bind_spmf p (\s1. map_spmf fst (cfunc_adv core1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (cfunc_adv core2 s2 ia))" + and sim_cfunc_adv: "\p q ia oa. \ S_core p q; ia \ IA \ \ + S_core (cond_spmf_fst (bind_spmf p (\s1. cfunc_adv core1 s1 ia)) oa) + (cond_spmf_fst (bind_spmf q (\s2. cfunc_adv core2 s2 ia)) oa)" + and step_cfunc_usr: "\p q iu. \ S_core p q; iu \ IU \ + \ bind_spmf p (\s1. map_spmf fst (cfunc_usr core1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (cfunc_usr core2 s2 iu))" + and sim_cfunc_usr: "\p q iu ou. \ S_core p q; iu \ IU \ \ + S_core (cond_spmf_fst (bind_spmf p (\s1. cfunc_usr core1 s1 iu)) ou) + (cond_spmf_fst (bind_spmf q (\s2. cfunc_usr core2 s2 iu)) ou)" + by(rule trace_core_eq_complete) blast + + show ?thesis using core_start + proof(coinduct rule: trace'_eqI_sim[consumes 1, case_names step sim]) + case (step p q a) + then consider (cpoke) e where "a = Inl e" "e \ E" + | (cfunc_adv) ia where "a = Inr (Inl ia)" "ia \ IA" + | (cfunc_usr) iu where "a = Inr (Inr iu)" "iu \ IU" by auto + then show ?case + proof cases + case cpoke + with step_cpoke[OF step(1), of e] show ?thesis + by(simp add: spmf.map_comp o_def map_spmf_const weight_bind_spmf) + (auto intro!: spmf_eqI simp add: spmf_bind spmf_scale_spmf max_def min_absorb2 weight_spmf_le_1) + next + case cfunc_adv + with step_cfunc_adv[OF step(1) cfunc_adv(2)] show ?thesis + by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric] map_bind_spmf[unfolded o_def, symmetric]) + next + case cfunc_usr + with step_cfunc_usr[OF step(1) cfunc_usr(2)] show ?thesis + by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric] map_bind_spmf[unfolded o_def, symmetric]) + qed + next + case (sim p q a res b s') + then consider (cpoke) e where "a = Inl e" "e \ E" + | (cfunc_adv) ia where "a = Inr (Inl ia)" "ia \ IA" + | (cfunc_usr) iu where "a = Inr (Inr iu)" "iu \ IU" by auto + then show ?case + proof cases + case cpoke + with sim_cpoke[OF sim(1) , of e] sim show ?thesis + by(clarsimp simp add: map_bind_spmf[unfolded o_def, symmetric]) + next + case cfunc_adv + with sim_cfunc_adv[OF sim(1) cfunc_adv(2)] sim show ?thesis + apply(clarsimp simp add: map_bind_spmf[unfolded o_def, symmetric] apfst_def map_prod_def) + apply(subst (1 2) cond_spmf_fst_map_prod_inj) + apply(simp_all add: o_def[symmetric] inj_compose del: o_apply) + done + next + case cfunc_usr + with sim_cfunc_usr[OF sim(1) cfunc_usr(2)] sim show ?thesis + apply(clarsimp simp add: map_bind_spmf[unfolded o_def, symmetric] apfst_def map_prod_def) + apply(subst (1 2) cond_spmf_fst_map_prod_inj) + apply(simp_all add: o_def[symmetric] inj_compose del: o_apply) + done + qed + qed +qed + +lemma trace_eq_callee_of_restI: + "trace_callee_eq (callee_of_rest rest1) (callee_of_rest rest2) (IA <+> IU) p q" + if "trace_rest_eq rest1 rest2 IA IU p q" +proof - + from that obtain S_rest + where rest_start: "S_rest p q" + and step_rfunc_adv: "\p q ia. \ S_rest p q; ia \ IA \ + \ bind_spmf p (\s1. map_spmf fst (rfunc_adv rest1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (rfunc_adv rest2 s2 ia))" + and sim_rfunc_adv: "\p q ia oa. \ S_rest p q; ia \ IA \ \ + S_rest (cond_spmf_fst (bind_spmf p (\s1. rfunc_adv rest1 s1 ia)) oa) + (cond_spmf_fst (bind_spmf q (\s2. rfunc_adv rest2 s2 ia)) oa)" + and step_rfunc_usr: "\p q iu. \ S_rest p q; iu \ IU \ + \ bind_spmf p (\s1. map_spmf fst (rfunc_usr rest1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (rfunc_usr rest2 s2 iu))" + and sim_rfunc_usr: "\p q iu ou. \ S_rest p q; iu \ IU \ \ + S_rest (cond_spmf_fst (bind_spmf p (\s1. rfunc_usr rest1 s1 iu)) ou) + (cond_spmf_fst (bind_spmf q (\s2. rfunc_usr rest2 s2 iu)) ou)" + by(rule trace_rest_eq_complete) blast + + show ?thesis using rest_start + proof(coinduct rule: trace'_eqI_sim[consumes 1, case_names step sim]) + case (step p q a) + then consider (rfunc_adv) ia where "a = Inl ia" "ia \ IA" + | (rfunc_usr) iu where "a = Inr iu" "iu \ IU" by auto + then show ?case + proof cases + case rfunc_adv + with step_rfunc_adv[OF step(1) rfunc_adv(2)] show ?thesis + by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric] map_bind_spmf[unfolded o_def, symmetric]) + next + case rfunc_usr + with step_rfunc_usr[OF step(1) rfunc_usr(2)] show ?thesis + by(simp add: spmf.map_comp)(simp add: spmf.map_comp[symmetric] map_bind_spmf[unfolded o_def, symmetric]) + qed + next + case (sim p q a res b s') + then consider (rfunc_adv) ia where "a = Inl ia" "ia \ IA" + | (rfunc_usr) iu where "a = Inr iu" "iu \ IU" by auto + then show ?case + proof cases + case rfunc_adv + with sim_rfunc_adv[OF sim(1) rfunc_adv(2)] sim show ?thesis + by(clarsimp simp add: map_bind_spmf[unfolded o_def, symmetric] apfst_def map_prod_def) + (subst (1 2) cond_spmf_fst_map_prod_inj; simp) + next + case rfunc_usr + with sim_rfunc_usr[OF sim(1) rfunc_usr(2)] sim show ?thesis + by(clarsimp simp add: map_bind_spmf[unfolded o_def, symmetric] apfst_def map_prod_def) + (subst (1 2) cond_spmf_fst_map_prod_inj; simp) + qed + qed +qed + +lemma trace_callee_resource_of_oracle: + "trace_callee run_resource (map_spmf (resource_of_oracle callee) p) = trace_callee callee p" + (is "?lhs = ?rhs") +proof(intro ext) + show "?lhs tr x = ?rhs tr x" for tr x + proof(induction tr arbitrary: p) + case Nil show ?case by(simp add: bind_map_spmf o_def spmf.map_comp) + next + case (Cons a tr) + obtain y z where a [simp]: "a = (y, z)" by(cases a) + have "trace_callee run_resource (map_spmf (RES callee) p) (a # tr) x = + trace_callee run_resource (cond_spmf_fst (map_spmf (\(x, y). (x, RES callee y)) (p \ (\x. (callee x y)))) z) tr x" + by(clarsimp simp add: bind_map_spmf o_def map_prod_def map_bind_spmf) + also have "\ = trace_callee run_resource (map_spmf (RES callee) (cond_spmf_fst (p \ (\x. (callee x y))) z)) tr x" + by(subst cond_spmf_fst_map_prod_inj) simp_all + finally show ?case using Cons.IH by simp + qed +qed + +lemma trace_callee_resource_of_oracle': + "trace_callee run_resource (return_spmf (resource_of_oracle callee s)) = trace_callee callee (return_spmf s)" + using trace_callee_resource_of_oracle[where p="return_spmf s"] + by simp + +lemma trace_eq_resource_of_oracle: + "trace_eq A (map_spmf (resource_of_oracle callee1) p) (map_spmf (resource_of_oracle callee2) q) = + trace_callee_eq callee1 callee2 A p q" + unfolding trace_callee_eq_def trace_callee_resource_of_oracle by(rule refl) + +lemma WT_fuse_converter [WT_intro]: + "(\AC \\<^sub>\ map_\ id fst \AR) \\<^sub>\ (\UC \\<^sub>\ map_\ id fst \UR), (\E \\<^sub>\ (\AC \\<^sub>\ \UC)) \\<^sub>\ (\AR \\<^sub>\ \UR) \\<^sub>C fuse_converter \" + if "\x. \(y, es)\responses_\ \AR x. set es \ outs_\ \E" "\x. \(y, es)\responses_\ \UR x. set es \ outs_\ \E" + unfolding fuse_converter_def using that + by(intro WT_converter_of_callee) + (fastforce simp add: stateless_callee_def image_image intro: rev_image_eqI intro!: WT_gpv_pauses split: if_split_asm)+ + +theorem fuse_trace_eq: + fixes core1 :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and core2 :: "('s_core', 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and rest1 :: "('s_rest, 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more1) rest_scheme" + and rest2 :: "('s_rest', 'event, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more2) rest_scheme" + assumes core: "trace_core_eq core1 core2 (outs_\ \E) (outs_\ \CA) (outs_\ \CU) (return_spmf s_core) (return_spmf s_core')" + and rest: "trace_rest_eq rest1 rest2 (outs_\ \RA) (outs_\ \RU) (return_spmf s_rest) (return_spmf s_rest')" + and IC1: "callee_invariant_on (callee_of_core core1) IC1 (\E \\<^sub>\ (\CA \\<^sub>\ \CU))" "IC1 s_core" + and IC2: "callee_invariant_on (callee_of_core core2) IC2 (\E \\<^sub>\ (\CA \\<^sub>\ \CU))" "IC2 s_core'" + and IR1: "callee_invariant_on (callee_of_rest rest1) IR1 (\RA \\<^sub>\ \RU)" "IR1 s_rest" + and IR2: "callee_invariant_on (callee_of_rest rest2) IR2 (\RA \\<^sub>\ \RU)" "IR2 s_rest'" + and E1 [WT_intro]: "\x. \(y, es)\responses_\ \RA x. set es \ outs_\ \E" + and E2 [WT_intro]: "\x. \(y, es)\responses_\ \RU x. set es \ outs_\ \E" + shows "trace_callee_eq (fused_resource.fuse core1 rest1) (fused_resource.fuse core2 rest2) + ((outs_\ \CA <+> outs_\ \RA) <+> (outs_\ \CU <+> outs_\ \RU)) (return_spmf (s_core, s_rest)) (return_spmf (s_core', s_rest'))" +proof - + let ?\C = "\E \\<^sub>\ (\CA \\<^sub>\ \CU)" + let ?\R = "\RA \\<^sub>\ \RU" + let ?\' = "?\C \\<^sub>\ ?\R" + let ?\ = "(\CA \\<^sub>\ map_\ id fst \RA) \\<^sub>\ (\CU \\<^sub>\ map_\ id fst \RU)" + + interpret fuse1: fused_resource core1 s1 for s1 . + interpret fuse2: fused_resource core2 s2 for s2 . + interpret IC1: callee_invariant_on "callee_of_core core1" IC1 ?\C by fact + interpret IC2: callee_invariant_on "callee_of_core core2" IC2 ?\C by fact + interpret IR1: callee_invariant_on "callee_of_rest rest1" IR1 ?\R by fact + interpret IR2: callee_invariant_on "callee_of_rest rest2" IR2 ?\R by fact + + from core have "outs_\ ?\C \\<^sub>C callee_of_core core1(s_core) \ callee_of_core core2(s_core')" + by(simp add: trace_eq_callee_of_coreI) + hence "outs_\ ?\C \\<^sub>R RES (callee_of_core core1) s_core \ RES (callee_of_core core2) s_core'" by simp + moreover have "outs_\ ?\R \\<^sub>C callee_of_rest rest1(s_rest) \ callee_of_rest rest2(s_rest')" using rest + by(simp add: trace_eq_callee_of_restI) + hence "outs_\ ?\R \\<^sub>R RES (callee_of_rest rest1) s_rest \ RES (callee_of_rest rest2) s_rest'" by simp + ultimately have "outs_\ ?\' \\<^sub>R + RES (callee_of_core core1) s_core \ RES (callee_of_rest rest1) s_rest \ + RES (callee_of_core core2) s_core' \ RES (callee_of_rest rest2) s_rest'" + by(simp add: trace_eq'_parallel_resource) + hence "outs_\ ?\ \\<^sub>R fuse_converter \ (RES (callee_of_core core1) s_core \ RES (callee_of_rest rest1) s_rest) \ + fuse_converter \ (RES (callee_of_core core2) s_core' \ RES (callee_of_rest rest2) s_rest')" + by(rule attach_trace_eq')(intro WT_intro IC1.WT_resource_of_oracle IC1 IC2.WT_resource_of_oracle IC2 IR1.WT_resource_of_oracle IR1 IR2.WT_resource_of_oracle IR2)+ + hence "trace_eq' (outs_\ ?\) (resource_of_oracle (fuse1.fuse rest1) (s_core, s_rest)) (resource_of_oracle (fuse2.fuse rest2) (s_core', s_rest'))" + unfolding fuse_converter by simp + then show ?thesis by simp +qed + + +inductive trace_eq_simcl :: "('s1 spmf \ 's2 spmf \ bool) \ 's1 spmf \ 's2 spmf \ bool" + for S where + base: "trace_eq_simcl S p q" if "S p q" for p q +| bind_nat: "trace_eq_simcl S (bind_spmf p f) (bind_spmf p g)" +if "\x :: nat. x \ set_spmf p \ S (f x) (g x)" + +lemma trace_eq_simcl_bindI [intro?]: "trace_eq_simcl S (bind_spmf p f) (bind_spmf p g)" + if "\x. x \ set_spmf p \ S (f x) (g x)" + by(subst (1 2) bind_spmf_to_nat_on[symmetric])(auto intro!: trace_eq_simcl.bind_nat simp add: that) + +lemma trace_eq_simcl_bind: "trace_eq_simcl S (bind_spmf p f) (bind_spmf p g)" + if *: "\x :: 'a. x \ set_spmf p \ trace_eq_simcl S (f x) (g x)" +proof - + obtain P :: "'a \ nat spmf" and F G where + **: "\x. x \ set_spmf p \ f x = bind_spmf (P x) (F x) \ g x = bind_spmf (P x) (G x) \ (\y\set_spmf (P x). S (F x y) (G x y))" + apply(atomize_elim) + apply(subst choice_iff[symmetric])+ + apply(fastforce dest!: * elim!: trace_eq_simcl.cases intro: exI[where x="return_spmf _"]) + done + have "bind_spmf p f = bind_spmf (bind_spmf p (\x. map_spmf (Pair x) (P x))) (\(x, y). F x y)" + by(simp add: bind_map_spmf o_def ** cong: bind_spmf_cong) + moreover have "bind_spmf p g = bind_spmf (bind_spmf p (\x. map_spmf (Pair x) (P x))) (\(x, y). G x y)" + by(simp add: bind_map_spmf o_def ** cong: bind_spmf_cong) + ultimately show ?thesis by(simp only:)(rule trace_eq_simcl_bindI; clarsimp simp add: **) +qed + +lemma trace_eq_simcl_bind1_scale: "trace_eq_simcl S (bind_spmf p f) (scale_spmf (weight_spmf p) q)" + if "\x\set_spmf p. trace_eq_simcl S (f x) q" +proof - + have "trace_eq_simcl S (bind_spmf p f) (bind_spmf p (\_. q))" + by(rule trace_eq_simcl_bind)(simp add: that) + thus ?thesis by(simp add: bind_spmf_const) +qed + +lemma trace_eq_simcl_bind1: "trace_eq_simcl S (bind_spmf p f) q" + if "\x\set_spmf p. trace_eq_simcl S (f x) q" "lossless_spmf p" + using trace_eq_simcl_bind1_scale[OF that(1)] that(2) by(simp add: lossless_weight_spmfD) + +lemma trace_eq_simcl_bind2_scale: "trace_eq_simcl S (scale_spmf (weight_spmf q) p) (bind_spmf q f)" + if "\x\set_spmf q. trace_eq_simcl S p (f x)" +proof - + have "trace_eq_simcl S (bind_spmf q (\_. p)) (bind_spmf q f)" + by(rule trace_eq_simcl_bind)(simp add: that) + thus ?thesis by(simp add: bind_spmf_const) +qed + +lemma trace_eq_simcl_bind2: "trace_eq_simcl S p (bind_spmf q f)" + if "\x\set_spmf q. trace_eq_simcl S p (f x)" "lossless_spmf q" + using trace_eq_simcl_bind2_scale[OF that(1)] that(2) by(simp add: lossless_weight_spmfD) + +lemma trace_eq_simcl_return_pmf_None [simp, intro!]: "trace_eq_simcl S (return_pmf None) (return_pmf None)" + for S :: "'s1 spmf \ 's2 spmf \ bool" +proof - + have "trace_eq_simcl S (bind_spmf (return_pmf None) (undefined :: nat \ 's1 spmf)) (bind_spmf (return_pmf None) (undefined :: nat \ 's2 spmf))" + by(rule trace_eq_simcl_bindI) simp + then show ?thesis by simp +qed + +lemma trace_eq_simcl_map: "trace_eq_simcl S (map_spmf f p) (map_spmf g p)" + if "\x\set_spmf p. S (return_spmf (f x)) (return_spmf (g x))" + unfolding map_spmf_conv_bind_spmf + by(rule trace_eq_simcl_bindI)(simp add: that) + +lemma trace_eq_simcl_map1: "trace_eq_simcl S (map_spmf f p) q" + if "\x\set_spmf p. trace_eq_simcl S (return_spmf (f x)) q" "lossless_spmf p" + unfolding map_spmf_conv_bind_spmf + by(rule trace_eq_simcl_bind1)(simp_all add: that) + +lemma trace_eq_simcl_map2: "trace_eq_simcl S p (map_spmf f q)" + if "\x\set_spmf q. trace_eq_simcl S p (return_spmf (f x))" "lossless_spmf q" + unfolding map_spmf_conv_bind_spmf + by(rule trace_eq_simcl_bind2)(simp_all add: that) + +lemma trace_eq_simcl_return_spmf [simp]: "trace_eq_simcl S (return_spmf x) (return_spmf y) \ S (return_spmf x) (return_spmf y)" + apply(rule iffI) + subgoal by(erule trace_eq_simcl.cases; clarsimp dest!: sym[where s="return_spmf _"])(auto 4 4 simp add: bind_eq_return_spmf dest!: lossless_spmfD_set_spmf_nonempty) + by(simp add: trace_eq_simcl.base) + +lemma trace_eq_simcl_callee: + fixes callee1 :: "('a, 'b, 's1) callee" and callee2 :: "('a, 'b, 's2) callee" + assumes step: "\p q a. \ S p q; a \ A \ \ + bind_spmf p (\s. map_spmf fst (callee1 s a)) = bind_spmf q (\s. map_spmf fst (callee2 s a))" + and sim: "\p q a res b s'. \ S p q; a \ A; res \ set_spmf q; (b, s') \ set_spmf (callee2 res a) \ + \ trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" + and start: "trace_eq_simcl S p q" and a: "a \ A" + shows trace_eq_simcl_callee_step: "bind_spmf p (\s. map_spmf fst (callee1 s a)) = bind_spmf q (\s. map_spmf fst (callee2 s a))" (is "?step") + and trace_eq_simcl_callee_sim: "\res b s'. \ res \ set_spmf q; (b, s') \ set_spmf (callee2 res a) \ + \ trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" (is "\res b s'. \ ?res res; ?b res b s' \ \ ?sim res b s'") +proof - + show eq: ?step using start a by cases(auto intro!: bind_spmf_cong intro: step) + show "?sim res b s'" if "?res res" "?b res b s'" for res b s' using start + proof cases + case base then show ?thesis using a that by(rule sim) + next + case (bind_nat X f g) + let ?Y = "cond_bind_spmf_fst X (\y. map_spmf fst (bind_spmf (f y) (\s. callee1 s a))) b" + let ?Y' = "cond_bind_spmf_fst X (\y. map_spmf fst (bind_spmf (g y) (\s. callee2 s a))) b" + have "cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b = bind_spmf ?Y (\x. cond_spmf_fst (bind_spmf (f x) (\s. callee1 s a)) b)" + unfolding bind_nat by(simp add: cond_spmf_fst_bind o_def) + moreover have "cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b = bind_spmf ?Y' (\x. cond_spmf_fst (bind_spmf (g x) (\s. callee2 s a)) b)" + unfolding bind_nat by(simp add: cond_spmf_fst_bind o_def) + moreover have "?Y = ?Y'" using bind_nat eq + by(intro spmf_eqI)(fastforce simp add: map_bind_spmf o_def spmf_eq_0_set_spmf dest: step[OF _ a]) + ultimately + show "trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" using bind_nat a + by(simp)(rule trace_eq_simcl_bind; auto intro!: sim simp add: bind_UNION) + qed +qed + +proposition trace'_eqI_sim_upto: + fixes callee1 :: "('a, 'b, 's1) callee" and callee2 :: "('a, 'b, 's2) callee" + assumes start: "S p q" + and step: "\p q a. \ S p q; a \ A \ \ + bind_spmf p (\s. map_spmf fst (callee1 s a)) = bind_spmf q (\s. map_spmf fst (callee2 s a))" + and sim: "\p q a res b s'. \ S p q; a \ A; res \ set_spmf q; (b, s') \ set_spmf (callee2 res a) \ + \ trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" + shows "trace_callee_eq callee1 callee2 A p q" +proof - + let ?S = "trace_eq_simcl S" + from start have "?S p q" by(rule trace_eq_simcl.base) + then show ?thesis by(rule trace'_eqI_sim)(rule trace_eq_simcl_callee[OF step sim]; assumption)+ +qed + +lemma trace_core_eq_simI_upto: + fixes core1 :: "('s_core, 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and core2 :: "('s_core', 'event, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and S :: "'s_core spmf \ 's_core' spmf \ bool" + assumes start: "S p q" + and step_cpoke: "\p q e. \ S p q; e \ E \ \ + weight_spmf (bind_spmf p (\s. cpoke core1 s e)) = weight_spmf (bind_spmf q (\s. cpoke core2 s e))" + and sim_cpoke: "\p q e. \ S p q; e \ E \ \ + trace_eq_simcl S (mk_lossless (bind_spmf p (\s. cpoke core1 s e))) (mk_lossless (bind_spmf q (\s. cpoke core2 s e)))" + and step_cfunc_adv: "\p q ia. \ S p q; ia \ IA \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_adv core1 s1 ia)) = bind_spmf q (\s2. map_spmf fst (cfunc_adv core2 s2 ia))" + and sim_cfunc_adv: "\p q ia s1 s2 s1' s2' oa. \ S p q; ia \ IA; + s1 \ set_spmf p; s2 \ set_spmf q; (oa, s1') \ set_spmf (cfunc_adv core1 s1 ia); (oa, s2') \ set_spmf (cfunc_adv core2 s2 ia) \ + \ trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s1. cfunc_adv core1 s1 ia)) oa) (cond_spmf_fst (bind_spmf q (\s2. cfunc_adv core2 s2 ia)) oa)" + and step_cfunc_usr: "\p q iu. \ S p q; iu \ IU \ \ + bind_spmf p (\s1. map_spmf fst (cfunc_usr core1 s1 iu)) = bind_spmf q (\s2. map_spmf fst (cfunc_usr core2 s2 iu))" + and sim_cfunc_usr: "\p q iu s1 s2 s1' s2' ou. \ S p q; iu \ IU; + s1 \ set_spmf p; s2 \ set_spmf q; (ou, s1') \ set_spmf (cfunc_usr core1 s1 iu); (ou, s2') \ set_spmf (cfunc_usr core2 s2 iu) \ + \ trace_eq_simcl S (cond_spmf_fst (bind_spmf p (\s1. cfunc_usr core1 s1 iu)) ou) (cond_spmf_fst (bind_spmf q (\s2. cfunc_usr core2 s2 iu)) ou)" + shows "trace_core_eq core1 core2 E IA IU p q" +proof - + let ?S = "trace_eq_simcl S" + from start have "?S p q" by(rule trace_eq_simcl.base) + then show ?thesis + proof(rule trace_core_eq_simI, goal_cases Step_cpoke Sim_cpoke Step_cfunc_adv Sim_cfunc_adv Step_cfunc_usr Sim_cfunc_usr) + { case (Step_cpoke p q e) + then show ?case using step_cpoke + by cases(auto simp add: weight_bind_spmf o_def intro!: Bochner_Integration.integral_cong_AE) } + note eq = this + + case (Sim_cpoke p q e) then show ?case + proof cases + case base then show ?thesis using Sim_cpoke(2) by(rule sim_cpoke) + next + case (bind_nat X f g) + then have "cond_bind_spmf X (\y. f y \ (\s. cpoke core1 s e)) UNIV = cond_bind_spmf X (\y. g y \ (\s. cpoke core2 s e)) UNIV" + using eq[OF Sim_cpoke] step_cpoke Sim_cpoke + by(intro spmf_eqI)(simp add: weight_spmf_def measure_spmf_zero_iff bind_UNION spmf_eq_0_set_spmf) + then show ?thesis using bind_nat Sim_cpoke sim_cpoke + by(auto simp add: cond_bind_spmf cond_spmf_UNIV[symmetric] simp del: cond_spmf_UNIV intro: trace_eq_simcl_bind) + qed + next + { case (Step_cfunc_adv p q ia) + then show ?case using step_cfunc_adv by cases(auto intro!: bind_spmf_cong) } + note eq = this + + case (Sim_cfunc_adv p q ia s1 s2 s1' s2' oa) then show ?case + proof cases + case base then show ?thesis using Sim_cfunc_adv(2-) by(rule sim_cfunc_adv) + next + case (bind_nat X f g) + then have "cond_bind_spmf_fst X (\y. map_spmf fst (f y \ (\s1. cfunc_adv core1 s1 ia))) oa = + cond_bind_spmf_fst X (\y. map_spmf fst (g y \ (\s2. cfunc_adv core2 s2 ia))) oa" + using eq[OF Sim_cfunc_adv(1,2)] + by(intro spmf_eqI)(fastforce simp add: map_bind_spmf o_def spmf_eq_0_set_spmf dest: step_cfunc_adv[OF _ Sim_cfunc_adv(2)]) + then show ?thesis using bind_nat(3-) Sim_cfunc_adv(1-2) + unfolding bind_nat(1,2) bind_spmf_assoc + apply(subst (1 2) cond_spmf_fst_bind) + apply(simp add: o_def) + apply(rule trace_eq_simcl_bind) + apply clarsimp + apply(frule step_cfunc_adv[OF bind_nat(3) Sim_cfunc_adv(2), THEN arg_cong[where f="set_spmf"], THEN equalityD2]) + apply(clarsimp simp add: o_def bind_UNION) + apply(drule subsetD) + apply fastforce + apply(auto intro: sim_cfunc_adv) + done + qed + next + { case (Step_cfunc_usr p q iu) + then show ?case using step_cfunc_usr by cases(auto intro!: bind_spmf_cong) } + note eq = this + + case (Sim_cfunc_usr p q iu s1 s2 s1' s2' ou) then show ?case + proof cases + case base then show ?thesis using Sim_cfunc_usr(2-) by(rule sim_cfunc_usr) + next + case (bind_nat X f g) + then have "cond_bind_spmf_fst X (\y. map_spmf fst (f y \ (\s1. cfunc_usr core1 s1 iu))) ou = + cond_bind_spmf_fst X (\y. map_spmf fst (g y \ (\s2. cfunc_usr core2 s2 iu))) ou" + using eq[OF Sim_cfunc_usr(1,2)] + by(intro spmf_eqI)(fastforce simp add: map_bind_spmf o_def spmf_eq_0_set_spmf dest: step_cfunc_usr[OF _ Sim_cfunc_usr(2)]) + then show ?thesis using bind_nat(3-) Sim_cfunc_usr(1-2) + unfolding bind_nat(1,2) bind_spmf_assoc + apply(subst (1 2) cond_spmf_fst_bind) + apply(simp add: o_def) + apply(rule trace_eq_simcl_bind) + apply clarsimp + apply(frule step_cfunc_usr[OF bind_nat(3) Sim_cfunc_usr(2), THEN arg_cong[where f="set_spmf"], THEN equalityD2]) + apply(clarsimp simp add: o_def bind_UNION) + apply(drule subsetD) + apply fastforce + apply(auto intro: sim_cfunc_usr) + done + qed + qed +qed + + + +context + fixes core :: "('s_core, 'event1 + 'event2, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and rest :: "('s_rest, 'event2, 'iadv_rest, 'iusr_rest, 'oadv_rest, 'ousr_rest, 'more) rest_scheme" +begin + +primcorec core_with_rest :: + "('s_core \ 's_rest, 'event1, 'iadv_core + 'iadv_rest, 'iusr_core + 'iusr_rest, 'oadv_core + 'oadv_rest, 'ousr_core + 'ousr_rest) core" + where + "cpoke core_with_rest = (\(s_core, s_rest) e. map_spmf (\s_core'. (s_core', s_rest)) (cpoke core s_core (Inl e)))" + | "cfunc_adv core_with_rest = (\(s_core, s_rest) iadv. case iadv of + Inl iadv_core \ map_spmf (\(oadv_core, s_core'). (Inl oadv_core, (s_core', s_rest))) (cfunc_adv core s_core iadv_core) + | Inr iadv_rest \ + bind_spmf (rfunc_adv rest s_rest iadv_rest) (\((oadv_rest, es), s_rest'). + map_spmf (\s_core'. (Inr oadv_rest, (s_core', s_rest'))) (foldl_spmf (cpoke core) (return_spmf s_core) (map Inr es))))" + | "cfunc_usr core_with_rest = (\(s_core, s_rest) iusr. case iusr of + Inl iusr_core \ map_spmf (\(ousr_core, s_core'). (Inl ousr_core, (s_core', s_rest))) (cfunc_usr core s_core iusr_core) + | Inr iusr_rest \ + bind_spmf (rfunc_usr rest s_rest iusr_rest) (\((ousr_rest, es), s_rest'). + map_spmf (\s_core'. (Inr ousr_rest, (s_core', s_rest'))) (foldl_spmf (cpoke core) (return_spmf s_core) (map Inr es))))" + +end + +lemma fuse_core_with_rest: + fixes core :: "('s_core, 'event1 + 'event2, 'iadv_core, 'iusr_core, 'oadv_core, 'ousr_core) core" + and rest1 :: "('s_rest1, 'event1, 'iadv_rest1, 'iusr_rest1, 'oadv_rest1, 'ousr_rest1, 'more1) rest_scheme" + and rest2 :: "('s_rest2, 'event2, 'iadv_rest2, 'iusr_rest2, 'oadv_rest2, 'ousr_rest2, 'more2) rest_scheme" + shows + "fused_resource.fuse core (parallel_rest rest1 rest2) (s_core, (s_rest1, s_rest2)) = + map_fun (map_sum (lsumr \ map_sum id swap_sum) (lsumr \ map_sum id swap_sum)) (map_spmf (map_prod (map_sum (map_sum id swap_sum \ rsuml) (map_sum id swap_sum \ rsuml)) (map_prod id prod.swap \ rprodl))) + (fused_resource.fuse (core_with_rest core rest2) rest1 ((s_core, s_rest2), s_rest1))" + apply(rule ext) + subgoal for x + apply(cases "(parallel_rest rest1 rest2, (s_core, (s_rest1, s_rest2)), x)" rule: fused_resource.fuse.cases) + apply(auto simp add: fused_resource.fuse.simps map_bind_spmf bind_map_spmf map_prod_def split_def o_def parallel_eoracle_def parallel_oracle_def split!: sum.split intro!: bind_spmf_cong) + apply(subst foldl_spmf_pair_left[simplified split_def]; simp add: map_fun_def o_def bind_map_spmf)+ + done + done + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Goodies.thy b/thys/Constructive_Cryptography_CM/Goodies.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Goodies.thy @@ -0,0 +1,44 @@ +theory Goodies + imports + Main +begin + +primrec + inits :: "'a list \ 'a list list" + where + "inits [] = []" + | "inits (x # xs) = [x] # map ((#) x) (inits xs)" + +definition + inits_self :: "'a list \ ('a list \ 'a) list" + where + "inits_self xs = zip ([] # inits xs) xs" + +lemma inits_map: "inits (map f xs) = map (map f) (inits xs)" + by(induction xs) simp_all + +lemma inits_append [simp]: "inits (xs @ ys) = inits xs @ map ((@) xs) (inits ys)" + by(induction xs) (simp_all) + +lemma inits_self_simps [simp]: + "inits_self [] = []" + "inits_self (x # xs) = ([], x) # map (apfst ((#) x)) (inits_self xs)" + by(simp_all add: inits_self_def apfst_def map_prod_def zip_map1[symmetric]) + +lemma inits_self_map: "inits_self (map f xs) = map (map_prod (map f) f) (inits_self xs)" + by(induction xs) (simp_all add: apfst_def prod.map_comp o_def) + +lemma in_set_inits_self: "(ys, z) \ set (inits_self xs) \ (\zs. xs = ys @ z # zs)" + by(induction xs arbitrary: ys z)(auto simp add: Cons_eq_append_conv apfst_def map_prod_def) + +lemma foldl_append: "foldl (\s e. s @ [e]) s l = s @ l" + by (induction l arbitrary: s) auto + +lemma foldl_insert: "foldl (\A x. insert (f x) A) A xs = A \ (f ` set xs)" + by(induction xs arbitrary: A) simp_all + +lemma foldl_concat_prodl: "foldl (\(l, r) x. (l @ g r x, f r x)) (l, r) xs = + (l @ concat (map (\(ys, x). g (foldl f r ys) x) (inits_self xs)), foldl f r xs)" + by(induction xs arbitrary: l r) (simp_all add: split_def o_def) + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/More_CC.thy b/thys/Constructive_Cryptography_CM/More_CC.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/More_CC.thy @@ -0,0 +1,2214 @@ +theory More_CC imports + Constructive_Cryptography.Constructive_Cryptography +begin + +section \Material for Isabelle library\ + +lemma eq_alt_conversep: "(=) = (BNF_Def.Grp UNIV id)\\" + by(simp add: Grp_def fun_eq_iff) + +parametric_constant + swap_parametric [transfer_rule]: prod.swap_def + +lemma Sigma_parametric [transfer_rule]: includes lifting_syntax shows + "(rel_set A ===> (A ===> rel_set B) ===> rel_set (rel_prod A B)) Sigma Sigma" + unfolding Sigma_def by transfer_prover + +lemma empty_eq_Plus [simp]: "{} = A <+> B \ A = {} \ B = {}" + by auto + +lemma insert_Inl_Plus [simp]: "insert (Inl x) (A <+> B) = insert x A <+> B" by auto + +lemma insert_Inr_Plus [simp]: "insert (Inr x) (A <+> B) = A <+> insert x B" by auto + +lemma map_sum_image_Plus [simp]: "map_sum f g ` (A <+> B) = f ` A <+> g ` B" + by(auto intro: rev_image_eqI) + +lemma Plus_subset_Plus_iff [simp]: "A <+> B \ C <+> D \ A \ C \ B \ D" + by auto + +lemma map_sum_eq_Inl_iff: "map_sum f g x = Inl y \ (\x'. x = Inl x' \ y = f x')" + by(cases x) auto + +lemma map_sum_eq_Inr_iff: "map_sum f g x = Inr y \ (\x'. x = Inr x' \ y = g x')" + by(cases x) auto + +lemma surj_map_sum: "surj (map_sum f g)" if "surj f" "surj g" + apply(safe; simp) + subgoal for x using that + by(cases x)(auto 4 3 intro: image_eqI[where x="Inl _"] image_eqI[where x="Inr _"]) + done + +lemma bij_map_sumI [simp]: "bij (map_sum f g)" if "bij f" "bij g" + using that by(clarsimp simp add: bij_def sum.inj_map surj_map_sum) + +lemma inv_map_sum [simp]: + "\ bij f; bij g \ \ inv_into UNIV (map_sum f g) = map_sum (inv_into UNIV f) (inv_into UNIV g)" + by(rule inj_imp_inv_eq)(simp_all add: sum.map_comp sum.inj_map bij_def surj_iff sum.map_id) + + +context conditionally_complete_lattice begin + +lemma admissible_le1I: + "ccpo.admissible lub ord (\x. f x \ y)" + if "cont lub ord Sup (\) f" + by(rule ccpo.admissibleI)(auto simp add: that[THEN contD] intro!: cSUP_least) + +lemma admissible_le1_mcont [cont_intro]: + "ccpo.admissible lub ord (\x. f x \ y)" if "mcont lub ord Sup (\) f" + using that by(simp add: admissible_le1I mcont_def) + +end + +lemma eq_alt_conversep2: "(=) = ((BNF_Def.Grp UNIV id)\\)\\" + by(auto simp add: Grp_def fun_eq_iff) + +lemma nn_integral_indicator_singleton1 [simp]: + assumes [measurable]: "{y} \ sets M" + shows "(\\<^sup>+x. indicator {y} x * f x \M) = emeasure M {y} * f y" + by(simp add: mult.commute) + +lemma nn_integral_indicator_singleton1' [simp]: + assumes "{y} \ sets M" + shows "(\\<^sup>+x. indicator {x} y * f x \M) = emeasure M {y} * f y" + by(subst nn_integral_indicator_singleton1[symmetric, OF assms])(rule nn_integral_cong; simp split: split_indicator) + +subsection \Probabilities\ + +lemma pmf_eq_1_iff: "pmf p x = 1 \ p = return_pmf x" (is "?lhs = ?rhs") +proof(rule iffI) + assume ?lhs + have "pmf p i = 0" if "x \ i" for i + proof(rule antisym) + have "pmf p i + 1 \ pmf p i + pmf p x" using \?lhs\ by simp + also have "\ = measure (measure_pmf p) {i, x}" using that + by(subst measure_pmf.finite_measure_eq_sum_singleton)(simp_all add: pmf.rep_eq) + also have "\ \ 1" by(rule measure_pmf.subprob_measure_le_1) + finally show "pmf p i \ 0" by simp + qed(rule pmf_nonneg) + then show ?rhs if ?lhs + by(intro pmf_eqI)(auto simp add: that split: split_indicator) +qed simp + +lemma measure_spmf_cong: "measure (measure_spmf p) A = measure (measure_spmf p) B" + if "A \ set_spmf p = B \ set_spmf p" +proof - + have "measure (measure_spmf p) A = measure (measure_spmf p) (A \ set_spmf p) + measure (measure_spmf p) (A - set_spmf p)" + by(subst measure_spmf.finite_measure_Union[symmetric])(auto intro: arg_cong2[where f=measure]) + also have "measure (measure_spmf p) (A - set_spmf p) = 0" by(simp add: measure_spmf_zero_iff) + also have "0 = measure (measure_spmf p) (B - set_spmf p)" by(simp add: measure_spmf_zero_iff) + also have "measure (measure_spmf p) (A \ set_spmf p) + \ = measure (measure_spmf p) B" + unfolding that by(subst measure_spmf.finite_measure_Union[symmetric])(auto intro: arg_cong2[where f=measure]) + finally show ?thesis . +qed + +definition weight_spmf' where "weight_spmf' = weight_spmf" +lemma weight_spmf'_parametric [transfer_rule]: "rel_fun (rel_spmf A) (=) weight_spmf' weight_spmf'" + unfolding weight_spmf'_def by(rule weight_spmf_parametric) + +lemma bind_spmf_to_nat_on: + "bind_spmf (map_spmf (to_nat_on (set_spmf p)) p) (\n. f (from_nat_into (set_spmf p) n)) = bind_spmf p f" + by(simp add: bind_map_spmf cong: bind_spmf_cong) + +lemma try_cond_spmf_fst: + "try_spmf (cond_spmf_fst p x) q = (if x \ fst ` set_spmf p then cond_spmf_fst p x else q)" + by (metis cond_spmf_fst_eq_return_None lossless_cond_spmf_fst try_spmf_lossless try_spmf_return_None) + +lemma measure_try_spmf: + "measure (measure_spmf (try_spmf p q)) A = measure (measure_spmf p) A + pmf p None * measure (measure_spmf q) A" +proof - + have "emeasure (measure_spmf (try_spmf p q)) A = emeasure (measure_spmf p) A + pmf p None * emeasure (measure_spmf q) A" + by(fold nn_integral_spmf)(simp add: spmf_try_spmf nn_integral_add ennreal_mult' nn_integral_cmult) + then show ?thesis by(simp add: measure_spmf.emeasure_eq_measure ennreal_mult'[symmetric] ennreal_plus[symmetric] del: ennreal_plus) +qed + +lemma rel_spmf_OO_trans_strong: + "\ rel_spmf R p q; rel_spmf S q r \ \ rel_spmf (R OO eq_onp (\x. x \ set_spmf q) OO S) p r" + by(auto intro: rel_spmf_OO_trans rel_spmf_reflI simp add: eq_onp_def) + +lemma mcont2mcont_spmf [cont_intro]: + "mcont lub ord Sup (\) (\p. spmf (f p) x)" + if "mcont lub ord lub_spmf (ord_spmf (=)) f" + using that unfolding mcont_def + apply safe + subgoal by(rule monotone2monotone, rule monotone_spmf; simp) + apply(rule contI) + apply(subst contD[where f=f and luba=lub]; simp) + apply(subst cont_spmf[THEN contD]) + apply(erule (2) chain_imageI[OF _ monotoneD]) + apply simp + apply(simp add: image_image) + done + +lemma ord_spmf_try_spmf2: "ord_spmf R p (try_spmf p q)" if "rel_spmf R p p" +proof - + have "ord_spmf R (bind_pmf p return_pmf) (try_spmf p q)" unfolding try_spmf_def + by(rule rel_pmf_bindI[where R="rel_option R"]) + (use that in \auto simp add: rel_pmf_return_pmf1 elim!: option.rel_cases\) + then show ?thesis by(simp add: bind_return_pmf') +qed + +lemma ord_spmf_lossless_spmfD1: + assumes "ord_spmf R p q" + and "lossless_spmf p" + shows "rel_spmf R p q" + by (metis (no_types, lifting) assms lossless_iff_set_pmf_None option.simps(11) ord_option.cases pmf.rel_mono_strong) + +lemma restrict_spmf_mono: + "ord_spmf (=) p q \ ord_spmf (=) (p \ A) (q \ A)" + by(auto simp add: restrict_spmf_def pmf.rel_map elim!: pmf.rel_mono_strong elim: ord_option.cases) + +lemma restrict_lub_spmf: + assumes chain: "Complete_Partial_Order.chain (ord_spmf (=)) Y" + shows "restrict_spmf (lub_spmf Y) A = lub_spmf ((\p. restrict_spmf p A) ` Y)" (is "?lhs = ?rhs") +proof(cases "Y = {}") + case Y: False + have chain': "Complete_Partial_Order.chain (ord_spmf (=)) ((\p. p \ A) ` Y)" + using chain by(rule chain_imageI)(auto intro: restrict_spmf_mono) + show ?thesis by(rule spmf_eqI)(simp add: spmf_lub_spmf[OF chain'] Y image_image spmf_restrict_spmf spmf_lub_spmf[OF chain]) +qed simp + +lemma mono2mono_restrict_spmf [THEN spmf.mono2mono]: + shows monotone_restrict_spmf: "monotone (ord_spmf (=)) (ord_spmf (=)) (\p. p \ A)" + by(rule monotoneI)(rule restrict_spmf_mono) + +lemma mcont2mcont_restrict_spmf [THEN spmf.mcont2mcont, cont_intro]: + shows mcont_restrict_spmf: "mcont lub_spmf (ord_spmf (=)) lub_spmf (ord_spmf (=)) (\p. restrict_spmf p A)" + using monotone_restrict_spmf by(rule mcontI)(simp add: contI restrict_lub_spmf) + +lemma ord_spmf_case_option: "ord_spmf R (case x of None \ a | Some y \ b y) (case x of None \ a' | Some y \ b' y)" + if "ord_spmf R a a'" "\y. ord_spmf R (b y) (b' y)" using that by(cases x) auto + +lemma ord_spmf_map_spmfI: "ord_spmf (=) (map_spmf f p) (map_spmf f q)" if "ord_spmf (=) p q" + using that by(auto simp add: pmf.rel_map elim!: pmf.rel_mono_strong ord_option.cases) + +subsubsection \Conditional probabilities\ + +lemma mk_lossless_cond_spmf [simp]: "mk_lossless (cond_spmf p A) = cond_spmf p A" + by(simp add: cond_spmf_alt) + +context + fixes p :: "'a pmf" + and f :: "'a \ 'b pmf" + and A :: "'b set" + and F :: "'a \ real" + defines "F \ \x. pmf p x * measure (measure_pmf (f x)) A / measure (measure_pmf (bind_pmf p f)) A" +begin + +definition cond_bind_pmf :: "'a pmf" where "cond_bind_pmf = embed_pmf F" + +lemma cond_bind_pmf_nonneg: "F x \ 0" + by(simp add: F_def) + +context assumes defined: "A \ (\x\set_pmf p. set_pmf (f x)) \ {}" begin + +private lemma nonzero: "measure (measure_pmf (bind_pmf p f)) A > 0" + using defined by(auto 4 3 intro: measure_pmf_posI) + +lemma cond_bind_pmf_prob: "(\\<^sup>+ x. F x \count_space UNIV) = 1" +proof - + have nonzero': "(\\<^sup>+ x. ennreal (pmf p x) * ennreal (measure_pmf.prob (f x) A) \count_space UNIV) \ 0" + using defined by(auto simp add: nn_integral_0_iff_AE AE_count_space pmf_eq_0_set_pmf measure_pmf_zero_iff) + have finite: "(\\<^sup>+ x. ennreal (pmf p x) * ennreal (measure_pmf.prob (f x) A) \count_space UNIV) < \" (is "?lhs < _") + proof(rule order.strict_trans1) + show "?lhs \ (\\<^sup>+ x. ennreal (pmf p x) * 1 \count_space UNIV)" + by(rule nn_integral_mono)(simp add: mult_left_le) + show "\ < \" by(simp add: nn_integral_pmf_eq_1) + qed + have "(\\<^sup>+ x. F x \count_space UNIV) = + (\\<^sup>+ x. ennreal (pmf p x * measure_pmf.prob (f x) A)) / emeasure (measure_pmf (bind_pmf p f)) A" + using nonzero unfolding F_def measure_pmf.emeasure_eq_measure + by(simp add: divide_ennreal[symmetric] divide_ennreal_def nn_integral_multc) + also have "\ = 1" unfolding emeasure_bind_pmf + by(simp add: measure_pmf.emeasure_eq_measure nn_integral_measure_pmf ennreal_mult' nonzero' finite) + finally show ?thesis . +qed + +lemma pmf_cond_bind_pmf: "pmf cond_bind_pmf x = F x" + unfolding cond_bind_pmf_def by(rule pmf_embed_pmf[OF cond_bind_pmf_nonneg cond_bind_pmf_prob]) + +lemma set_cond_bind_pmf: "set_pmf cond_bind_pmf = {x\set_pmf p. set_pmf (f x) \ A \ {}}" + unfolding cond_bind_pmf_def + by(subst set_embed_pmf[OF cond_bind_pmf_nonneg cond_bind_pmf_prob]) + (auto simp add: F_def pmf_eq_0_set_pmf measure_pmf_zero_iff) + +lemma cond_bind_pmf: "cond_pmf (bind_pmf p f) A = bind_pmf cond_bind_pmf (\x. cond_pmf (f x) A)" + (is "?lhs = ?rhs") +proof(rule pmf_eqI) + fix i + have "ennreal (pmf ?lhs i) = ennreal (pmf ?rhs i)" + proof(cases "i \ A") + case True + have "ennreal (pmf ?lhs i) = (\\<^sup>+ x. ennreal (pmf p x) * ennreal (pmf (f x) i) / ennreal (measure_pmf.prob (p \ f) A) \count_space UNIV)" + using True defined + by(simp add: pmf_cond bind_UNION Int_commute divide_ennreal[symmetric] nonzero ennreal_pmf_bind) + (simp add: divide_ennreal_def nn_integral_multc[symmetric] nn_integral_measure_pmf) + also have "\ = (\\<^sup>+ x. ennreal (F x) * ennreal (pmf (cond_pmf (f x) A) i) \count_space UNIV)" + using True nonzero + apply(intro nn_integral_cong) + subgoal for x + by(clarsimp simp add: F_def ennreal_mult'[symmetric] divide_ennreal) + (cases "measure_pmf.prob (f x) A = 0"; auto simp add: pmf_cond pmf_eq_0_set_pmf measure_pmf_zero_iff) + done + also have "\ = ennreal (pmf ?rhs i)" + by(simp add: ennreal_pmf_bind nn_integral_measure_pmf pmf_cond_bind_pmf) + finally show ?thesis . + next + case False + then show ?thesis using defined + by(simp add: pmf_cond bind_UNION Int_commute pmf_eq_0_set_pmf set_cond_bind_pmf) + qed + then show "pmf ?lhs i = pmf ?rhs i" by simp +qed + +end + +end + + +lemma cond_spmf_try1: + "cond_spmf (try_spmf p q) A = cond_spmf p A" if "set_spmf q \ A = {}" + apply(rule spmf_eqI) + using that + apply(auto simp add: spmf_try_spmf measure_try_spmf measure_spmf_zero_iff) + apply(subst (2) spmf_eq_0_set_spmf[THEN iffD2]) + apply blast + apply simp + apply(simp add: measure_try_spmf measure_spmf_zero_iff) + done + +lemma cond_spmf_cong: "cond_spmf p A = cond_spmf p B" if "A \ set_spmf p = B \ set_spmf p" + apply(rule spmf_eqI) + using that by(auto simp add: measure_spmf_zero_iff spmf_eq_0_set_spmf measure_spmf_cong[OF that]) + +lemma cond_spmf_pair_spmf: + "cond_spmf (pair_spmf p q) (A \ B) = pair_spmf (cond_spmf p A) (cond_spmf q B)" (is "?lhs = ?rhs") +proof(rule spmf_eqI) + show "spmf ?lhs i = spmf ?rhs i" for i + proof(cases i) + case i [simp]: (Pair a b) + then show ?thesis by(simp add: measure_pair_spmf_times) + qed +qed + +lemma cond_spmf_pair_spmf1: + "cond_spmf_fst (map_spmf (\((x, s'), y). (f x, s', y)) (pair_spmf p q)) x = + pair_spmf (cond_spmf_fst (map_spmf (\(x, s'). (f x, s')) p) x) q" (is "?lhs = ?rhs") + if "lossless_spmf q" +proof - + have "?lhs = map_spmf (\((_, s'), y). (s', y)) (cond_spmf (pair_spmf p q) ((\((x, s'), y). (f x, s', y)) -` ({x} \ UNIV)))" + by(simp add: cond_spmf_fst_def spmf.map_comp o_def split_def) + also have "((\((x, s'), y). (f x, s', y)) -` ({x} \ UNIV)) = ((\(x, y). (f x, y)) -` ({x} \ UNIV)) \ UNIV" + by(auto) + also have "map_spmf (\((_, s'), y). (s', y)) (cond_spmf (pair_spmf p q) \) = ?rhs" + by(simp add: cond_spmf_fst_def cond_spmf_pair_spmf that spmf.map_comp pair_map_spmf1 apfst_def map_prod_def split_def) + finally show ?thesis . +qed + +lemma try_cond_spmf: "try_spmf (cond_spmf p A) q = (if set_spmf p \ A \ {} then cond_spmf p A else q)" + apply(clarsimp simp add: cond_spmf_def lossless_iff_set_pmf_None intro!: try_spmf_lossless) + apply(subst (asm) set_cond_pmf) + apply(auto simp add: in_set_spmf) + done + +lemma cond_spmf_try2: + "cond_spmf (try_spmf p q) A = (if lossless_spmf p then return_pmf None else cond_spmf q A)" if "set_spmf p \ A = {}" + apply(rule spmf_eqI) + using that + apply(auto simp add: spmf_try_spmf measure_try_spmf measure_spmf_zero_iff lossless_iff_pmf_None) + apply(subst spmf_eq_0_set_spmf[THEN iffD2]) + apply blast + apply(simp add: measure_spmf_zero_iff[THEN iffD2]) + done + + +definition cond_bind_spmf :: "'a spmf \ ('a \ 'b spmf) \ 'b set \ 'a spmf" where + "cond_bind_spmf p f A = + (if \x\set_spmf p. set_spmf (f x) \ A \ {} then + cond_bind_pmf p (\x. case x of None \ return_pmf None | Some x \ f x) (Some ` A) + else return_pmf None)" + +context begin + +private lemma defined: "\ y \ set_spmf (f x); y \ A; x \ set_spmf p \ + \ Some ` A \ (\x\set_pmf p. set_pmf (case x of None \ return_pmf None | Some x \ f x)) \ {}" + by(fastforce simp add: in_set_spmf bind_spmf_def) + +lemma spmf_cond_bind_spmf [simp]: + "spmf (cond_bind_spmf p f A) x = spmf p x * measure (measure_spmf (f x)) A / measure (measure_spmf (bind_spmf p f)) A" + by(clarsimp simp add: cond_bind_spmf_def measure_spmf_zero_iff bind_UNION pmf_cond_bind_pmf defined split!: if_split) + (fastforce simp add: in_set_spmf bind_spmf_def measure_measure_spmf_conv_measure_pmf)+ + +lemma set_cond_bind_spmf [simp]: + "set_spmf (cond_bind_spmf p f A) = {x\set_spmf p. set_spmf (f x) \ A \ {}}" + by(clarsimp simp add: cond_bind_spmf_def set_spmf_def bind_UNION) + (subst set_cond_bind_pmf; fastforce simp add: measure_measure_spmf_conv_measure_pmf) + +lemma cond_bind_spmf: "cond_spmf (bind_spmf p f) A = bind_spmf (cond_bind_spmf p f A) (\x. cond_spmf (f x) A)" + by(auto simp add: cond_spmf_def bind_UNION cond_bind_spmf_def split!: if_splits) + (fastforce split: option.splits simp add: cond_bind_pmf set_cond_bind_pmf defined in_set_spmf bind_spmf_def intro!: bind_pmf_cong[OF refl]) + +end + +lemma cond_spmf_fst_parametric [transfer_rule]: includes lifting_syntax shows + "(rel_spmf (rel_prod (=) B) ===> (=) ===> rel_spmf B) cond_spmf_fst cond_spmf_fst" + apply(rule rel_funI)+ + apply(clarsimp simp add: cond_spmf_fst_def spmf_rel_map elim!: rel_spmfE) + subgoal for x pq + by(subst (1 2) cond_spmf_cong[where B="fst -` ({x} \ UNIV) \ snd -` ({x} \ UNIV)"]) + (fastforce intro: rel_spmf_reflI)+ + done + +lemma cond_spmf_fst_map_prod: + "cond_spmf_fst (map_spmf (\(x, y). (f x, g x y)) p) (f x) = map_spmf (g x) (cond_spmf_fst p x)" + if "inj_on f (insert x (fst ` set_spmf p))" +proof - + have "cond_spmf p ((\(x, y). (f x, g x y)) -` ({f x} \ UNIV)) = cond_spmf p (((\(x, y). (f x, g x y)) -` ({f x} \ UNIV)) \ set_spmf p)" + by(rule cond_spmf_cong) simp + also have "((\(x, y). (f x, g x y)) -` ({f x} \ UNIV)) \ set_spmf p = ({x} \ UNIV) \ set_spmf p" + using that by(auto 4 3 dest: inj_onD intro: rev_image_eqI) + also have "cond_spmf p \ = cond_spmf p ({x} \ UNIV)" + by(rule cond_spmf_cong) simp + finally show ?thesis + by(auto simp add: cond_spmf_fst_def spmf.map_comp o_def split_def intro: map_spmf_cong) +qed + +lemma cond_spmf_fst_map_prod_inj: + "cond_spmf_fst (map_spmf (\(x, y). (f x, g x y)) p) (f x) = map_spmf (g x) (cond_spmf_fst p x)" + if "inj f" + apply(rule cond_spmf_fst_map_prod) + using that by(simp add: inj_on_def) + +definition cond_bind_spmf_fst :: "'a spmf \ ('a \ 'b spmf) \ 'b \ 'a spmf" where + "cond_bind_spmf_fst p f x = cond_bind_spmf p (map_spmf (\b. (b, ())) \ f) ({x} \ UNIV)" + +lemma cond_bind_spmf_fst_map_spmf_fst: + "cond_bind_spmf_fst p (map_spmf fst \ f) x = cond_bind_spmf p f ({x} \ UNIV)" (is "?lhs = ?rhs") +proof - + have [simp]: "(\x. (fst x, ())) -` ({x} \ UNIV) = {x} \ UNIV" by auto + have "?lhs = cond_bind_spmf p (\x. map_spmf (\x. (fst x, ())) (f x)) ({x} \ UNIV)" + by(simp add: cond_bind_spmf_fst_def spmf.map_comp o_def) + also have "\ = ?rhs" by(rule spmf_eqI)(simp add: measure_map_spmf map_bind_spmf[unfolded o_def, symmetric]) + finally show ?thesis . +qed + +lemma cond_spmf_fst_bind: "cond_spmf_fst (bind_spmf p f) x = + bind_spmf (cond_bind_spmf_fst p (map_spmf fst \ f) x) (\y. cond_spmf_fst (f y) x)" + by(simp add: cond_spmf_fst_def cond_bind_spmf map_bind_spmf cond_bind_spmf_fst_map_spmf_fst)(simp add: o_def) + +lemma spmf_cond_bind_spmf_fst [simp]: + "spmf (cond_bind_spmf_fst p f x) i = spmf p i * spmf (f i) x / spmf (bind_spmf p f) x" + by(simp add: cond_bind_spmf_fst_def) + (auto simp add: spmf_conv_measure_spmf measure_map_spmf map_bind_spmf[symmetric] intro!: arg_cong2[where f="(/)"] arg_cong2[where f="(*)"] arg_cong2[where f="measure"]) + +lemma set_cond_bind_spmf_fst [simp]: + "set_spmf (cond_bind_spmf_fst p f x) = {y \ set_spmf p. x \ set_spmf (f y)}" + by(auto simp add: cond_bind_spmf_fst_def intro: rev_image_eqI) + +lemma map_cond_spmf_fst: "map_spmf f (cond_spmf_fst p x) = cond_spmf_fst (map_spmf (apsnd f) p) x" + by(auto simp add: cond_spmf_fst_def spmf.map_comp intro!: map_spmf_cong arg_cong2[where f="cond_spmf"]) + +lemma cond_spmf_fst_try1: + "cond_spmf_fst (try_spmf p q) x = cond_spmf_fst p x" if "x \ fst ` set_spmf q" + using that + apply(simp add: cond_spmf_fst_def) + apply(subst cond_spmf_try1) + apply(auto intro: rev_image_eqI) + done + +lemma cond_spmf_fst_try2: + "cond_spmf_fst (try_spmf p q) x = (if lossless_spmf p then return_pmf None else cond_spmf_fst q x)" if "x \ fst ` set_spmf p" + using that + apply(simp add: cond_spmf_fst_def split!: if_split) + apply (metis cond_spmf_fst_def cond_spmf_fst_eq_return_None) + by (metis cond_spmf_fst_def cond_spmf_try2 lossless_cond_spmf lossless_cond_spmf_fst lossless_map_spmf) + +lemma cond_spmf_fst_map_inj: + "cond_spmf_fst (map_spmf (apfst f) p) (f x) = cond_spmf_fst p x" if "inj f" + by(auto simp add: cond_spmf_fst_def spmf.map_comp intro!: map_spmf_cong arg_cong2[where f=cond_spmf] dest: injD[OF that]) + +lemma cond_spmf_fst_pair_spmf1: + "cond_spmf_fst (map_spmf (\(x, y). (f x, g x y)) (pair_spmf p q)) a = + bind_spmf (cond_spmf_fst (map_spmf (\x. (f x, x)) p) a) (\x. map_spmf (g x) (mk_lossless q))" (is "?lhs = ?rhs") +proof - + have "(\(x, y). (f x, g x y)) -` ({a} \ UNIV) = f -` {a} \ UNIV" by(auto) + moreover have "(\x. (f x, x)) -` ({a} \ UNIV) = f -` {a}" by auto + ultimately show ?thesis + by(simp add: cond_spmf_fst_def spmf.map_comp o_def split_beta cond_spmf_pair_spmf) + (simp add: pair_spmf_alt_def map_bind_spmf o_def map_spmf_conv_bind_spmf) +qed + +lemma cond_spmf_fst_return_spmf': + "cond_spmf_fst (return_spmf (x, y)) z = (if x = z then return_spmf y else return_pmf None)" + by(simp add: cond_spmf_fst_def) + + + + + + +section \Material for CryptHOL\ + +lemma left_gpv_lift_spmf [simp]: "left_gpv (lift_spmf p) = lift_spmf p" + by(rule gpv.expand)(simp add: spmf.map_comp o_def) + +lemma right_gpv_lift_spmf [simp]: "right_gpv (lift_spmf p) = lift_spmf p" + by(rule gpv.expand)(simp add: spmf.map_comp o_def) + +lemma map'_lift_spmf: "map_gpv' f g h (lift_spmf p) = lift_spmf (map_spmf f p)" + by(rule gpv.expand)(simp add: gpv.map_sel spmf.map_comp o_def) + +lemma in_set_sample_uniform [simp]: "x \ set_spmf (sample_uniform n) \ x < n" + by(simp add: sample_uniform_def) + +lemma (in cyclic_group) inj_on_generator_iff [simp]: "\ x < order G; y < order G \ \ \<^bold>g [^] x = \<^bold>g [^] y \ x = y" + using inj_on_generator by(auto simp add: inj_on_def) + +lemma map_\_bot [simp]: "map_\ f g \ = \" + unfolding bot_\_def map_\_\_uniform by simp + +lemma map_\_Inr_plus [simp]: "map_\ Inr f (\1 \\<^sub>\ \2) = map_\ id (f \ Inr) \2" + by(rule \_eqI) auto + +lemma interaction_bound_map_gpv'_le: + defines "ib \ interaction_bound" + shows "interaction_bound consider (map_gpv' f g h gpv) \ ib (consider \ g) gpv" +proof(induction arbitrary: gpv rule: interaction_bound_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step interaction_bound') + show ?case unfolding ib_def + by(subst interaction_bound.simps) + (auto simp add: image_comp ib_def split: generat.split intro!: SUP_mono rev_bexI step.IH[unfolded ib_def]) +qed + +lemma interaction_bounded_by_map_gpv' [interaction_bound]: + assumes "interaction_bounded_by (consider \ g) gpv n" + shows "interaction_bounded_by consider (map_gpv' f g h gpv) n" + using assms interaction_bound_map_gpv'_le[of "consider" f g h gpv] by(simp add: interaction_bounded_by.simps) + +lemma map_gpv'_bind_gpv: + "map_gpv' f g h (bind_gpv gpv F) = bind_gpv (map_gpv' id g h gpv) (\x. map_gpv' f g h (F x))" + by(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + (auto simp del: bind_gpv_sel' simp add: bind_gpv.sel spmf_rel_map bind_map_spmf generat.rel_map rel_fun_def intro!: rel_spmf_bind_reflI rel_spmf_reflI generat.rel_refl_strong split!: generat.split) + +lemma exec_gpv_map_gpv': + "exec_gpv callee (map_gpv' f g h gpv) s = + map_spmf (map_prod f id) (exec_gpv (map_fun id (map_fun g (map_spmf (map_prod h id))) callee) gpv s)" + using exec_gpv_parametric'[ + where S="(=)" and CALL="BNF_Def.Grp UNIV g" and R="conversep (BNF_Def.Grp UNIV h)" and A="BNF_Def.Grp UNIV f", + unfolded rel_gpv''_Grp, simplified] + apply(subst (asm) (2) conversep_eq[symmetric]) + apply(subst (asm) prod.rel_conversep) + apply(subst (asm) (2 4) eq_alt) + apply(subst (asm) prod.rel_Grp) + apply simp + apply(subst (asm) spmf_rel_conversep) + apply(subst (asm) option.rel_Grp) + apply(subst (asm) pmf.rel_Grp) + apply simp + apply(subst (asm) prod.rel_Grp) + apply simp + apply(subst (asm) (1 3) conversep_conversep[symmetric]) + apply(subst (asm) rel_fun_conversep) + apply(subst (asm) rel_fun_Grp) + apply(subst (asm) rel_fun_conversep) + apply simp + apply(simp add: option.rel_Grp pmf.rel_Grp fun.rel_Grp) + apply(simp add: rel_fun_def BNF_Def.Grp_def o_def map_fun_def) + apply(erule allE)+ + apply(drule fun_cong) + apply(erule trans) + apply simp + done + +lemma colossless_gpv_sub_gpvs: + assumes "colossless_gpv \ gpv" "gpv' \ sub_gpvs \ gpv" + shows "colossless_gpv \ gpv'" +using assms(2,1) by(induction)(auto dest: colossless_gpvD) + +lemma pfinite_gpv_sub_gpvs: + assumes "pfinite_gpv \ gpv" "gpv' \ sub_gpvs \ gpv" "\ \g gpv \" + shows "pfinite_gpv \ gpv'" + using assms(2,1,3) by(induction)(auto dest: pfinite_gpv_ContD WT_gpvD) + +lemma pfinite_gpv_id_oracle [simp]: "pfinite_gpv \ (id_oracle s x)" if "x \ outs_\ \" + by(simp add: id_oracle_def pgen_lossless_gpv_PauseI[OF that]) + +subsection \@{term try_gpv}\ + +lemma plossless_gpv_try_gpvI: + assumes "pfinite_gpv \ gpv" + and "\ colossless_gpv \ gpv \ plossless_gpv \ gpv'" + shows "plossless_gpv \ (TRY gpv ELSE gpv')" + using assms unfolding pgen_lossless_gpv_def + by(cases "colossless_gpv \ gpv")(simp cong: expectation_gpv_cong_fail, simp) + +lemma WT_gpv_try_gpvI [WT_intro]: + assumes "\ \g gpv \" + and "\ colossless_gpv \ gpv \ \ \g gpv' \" + shows "\ \g try_gpv gpv gpv' \" + using assms by(coinduction arbitrary: gpv)(auto 4 4 dest: WT_gpvD colossless_gpvD split: if_split_asm) + +lemma (in callee_invariant_on) exec_gpv_try_gpv: + fixes exec_gpv1 + defines "exec_gpv1 \ exec_gpv" + assumes WT: "\ \g gpv \" + and pfinite: "pfinite_gpv \ gpv" + and I: "I s" + and f: "\s. I s \ f (x, s) = z" + and lossless: "\s x. \x \ outs_\ \; I s\ \ lossless_spmf (callee s x)" + shows "map_spmf f (exec_gpv callee (try_gpv gpv (Done x)) s) = + try_spmf (map_spmf f (exec_gpv1 callee gpv s)) (return_spmf z)" + (is "?lhs = ?rhs") +proof - + note [[show_variants]] + have le: "ord_spmf (=) ?lhs ?rhs" using WT I + proof(induction arbitrary: gpv s rule: exec_gpv_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step exec_gpv') + show ?case using step.prems unfolding exec_gpv1_def + apply(subst exec_gpv.simps) + apply(simp add: map_spmf_bind_spmf) + apply(subst (1 2) try_spmf_def) + apply(simp add: map_bind_pmf bind_spmf_pmf_assoc o_def) + apply(simp add: bind_spmf_def bind_map_pmf bind_assoc_pmf) + apply(rule rel_pmf_bindI[where R="eq_onp (\x. x \ set_pmf (the_gpv gpv))"]) + apply(rule pmf.rel_refl_strong) + apply(simp add: eq_onp_def) + apply(clarsimp split!: option.split generat.split simp add: bind_return_pmf f map_spmf_bind_spmf o_def eq_onp_def) + apply(simp add: bind_spmf_def bind_assoc_pmf) + subgoal for out c + apply(rule rel_pmf_bindI[where R="eq_onp (\x. x \ set_pmf (callee s out))"]) + apply(rule pmf.rel_refl_strong) + apply(simp add: eq_onp_def) + apply(clarsimp split!: option.split simp add: eq_onp_def) + apply(simp add: in_set_spmf[symmetric]) + apply(rule spmf.leq_trans) + apply(rule step.IH) + apply(frule (1) WT_gpvD) + apply(erule (1) WT_gpvD) + apply(drule WT_callee) + apply(erule (2) WT_calleeD) + apply(frule (1) WT_gpvD) + apply(erule (2) callee_invariant) + apply(simp add: try_spmf_def exec_gpv1_def) + done + done + qed + + have "lossless_spmf ?lhs" + apply simp + apply(rule plossless_exec_gpv) + apply(rule plossless_gpv_try_gpvI) + apply(rule pfinite) + apply simp + apply(rule WT_gpv_try_gpvI) + apply(simp add: WT) + apply simp + apply(simp add: lossless) + apply(simp add: I) + done + from ord_spmf_lossless_spmfD1[OF le this] show ?thesis by(simp add: spmf_rel_eq) +qed + +lemma try_gpv_bind_gen_lossless': \ \generalises @{thm try_gpv_bind_gen_lossless}\ + assumes lossless: "gen_lossless_gpv b \ gpv" + and WT1: "\ \g gpv \" + and WT2: "\ \g gpv' \" + and WTf: "\x. x \ results_gpv \ gpv \ \ \g f x \" + shows "eq_\_gpv (=) \ (TRY bind_gpv gpv f ELSE gpv') (bind_gpv gpv (\x. TRY f x ELSE gpv'))" + using lossless WT1 WTf +proof(coinduction arbitrary: gpv) + case (eq_\_gpv gpv) + note [simp] = spmf_rel_map generat.rel_map map_spmf_bind_spmf + and [intro!] = rel_spmf_reflI rel_generat_reflI rel_funI + show ?case using gen_lossless_gpvD[OF eq_\_gpv(1)] WT_gpvD[OF eq_\_gpv(2)] WT_gpvD[OF WT2] WT_gpvD[OF eq_\_gpv(3)[rule_format, OF results_gpv.Pure]] WT2 + apply(auto simp del: bind_gpv_sel' simp add: bind_gpv.sel try_spmf_bind_spmf_lossless generat.map_comp o_def intro!: rel_spmf_bind_reflI rel_spmf_try_spmf split!: generat.split) + apply(auto 4 4 intro!: eq_\_gpv(3)[rule_format] eq_\_gpv_reflI eq_\_generat_reflI intro: results_gpv.IO WT_intro) + done +qed + +\ \We instantiate the parameter @{term b} such that it can be used as a conditional simp rule.\ +lemmas try_gpv_bind_lossless' = try_gpv_bind_gen_lossless'[where b=False] + and try_gpv_bind_colossless' = try_gpv_bind_gen_lossless'[where b=True] + +lemma try_gpv_bind_gpv: + "try_gpv (bind_gpv gpv f) gpv' = + bind_gpv (try_gpv (map_gpv Some id gpv) (Done None)) (\x. case x of None \ gpv' | Some x' \ try_gpv (f x') gpv')" + by(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + (auto simp add: rel_fun_def generat.rel_map bind_return_pmf spmf_rel_map map_bind_spmf o_def bind_gpv.sel bind_map_spmf try_spmf_def bind_spmf_def spmf.map_comp bind_map_pmf bind_assoc_pmf gpv.map_sel simp del: bind_gpv_sel' intro!: rel_pmf_bind_reflI generat.rel_refl_strong rel_spmf_reflI split!: option.split generat.split) + +lemma bind_gpv_try_gpv_map_Some: + "bind_gpv (try_gpv (map_gpv Some id gpv) (Done None)) (\x. case x of None \ Fail | Some y \ f y) = + bind_gpv gpv f" + by(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + (auto simp add: bind_gpv.sel map_bind_spmf bind_map_spmf try_spmf_def bind_spmf_def spmf_rel_map bind_map_pmf gpv.map_sel bind_assoc_pmf bind_return_pmf generat.rel_map rel_fun_def simp del: bind_gpv_sel' intro!: rel_pmf_bind_reflI rel_spmf_reflI generat.rel_refl_strong split!: option.split generat.split) + +lemma try_gpv_left_gpv: + assumes "\ \g gpv \" and WT2: "\ \g gpv' \" + shows "eq_\_gpv (=) (\ \\<^sub>\ \') (try_gpv (left_gpv gpv) (left_gpv gpv')) (left_gpv (try_gpv gpv gpv'))" + using assms(1) + apply(coinduction arbitrary: gpv) + apply(auto simp add: map_try_spmf spmf.map_comp o_def generat.map_comp spmf_rel_map intro!: rel_spmf_try_spmf rel_spmf_reflI) + subgoal for gpv generat by(cases generat)(auto dest: WT_gpvD) + subgoal for gpv generat using WT2 + by(cases generat)(auto 4 4 dest: WT_gpvD intro!: eq_\_gpv_reflI WT_gpv_left_gpv) + done + +lemma try_gpv_right_gpv: + assumes "\' \g gpv \" and WT2: "\' \g gpv' \" + shows "eq_\_gpv (=) (\ \\<^sub>\ \') (try_gpv (right_gpv gpv) (right_gpv gpv')) (right_gpv (try_gpv gpv gpv'))" + using assms(1) + apply(coinduction arbitrary: gpv) + apply(auto simp add: map_try_spmf spmf.map_comp o_def generat.map_comp spmf_rel_map intro!: rel_spmf_try_spmf rel_spmf_reflI) + subgoal for gpv generat by(cases generat)(auto dest: WT_gpvD) + subgoal for gpv generat using WT2 + by(cases generat)(auto 4 4 dest: WT_gpvD intro!: eq_\_gpv_reflI WT_gpv_right_gpv) + done + +lemma bind_try_Done_Fail: "bind_gpv (TRY gpv ELSE Done x) f = bind_gpv gpv f" if "f x = Fail" + apply(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + apply(auto simp del: bind_gpv_sel' simp add: bind_gpv.sel map_bind_spmf bind_map_spmf try_spmf_def bind_spmf_def map_bind_pmf bind_assoc_pmf bind_map_pmf bind_return_pmf spmf.map_comp o_def that rel_fun_def intro!: rel_pmf_bind_reflI rel_spmf_reflI generat.rel_refl_strong split!: option.split generat.split) + done + + +lemma inline_map_gpv': + "inline callee (map_gpv' f g h gpv) s = + map_gpv (apfst f) id (inline (map_fun id (map_fun g (map_gpv (apfst h) id)) callee) gpv s)" + using inline_parametric'[where S="(=)" and C="BNF_Def.Grp UNIV g" and R="conversep (BNF_Def.Grp UNIV h)" and A="BNF_Def.Grp UNIV f" and C'="(=)" and R'="(=)"] + apply(subst (asm) (2 3 8) eq_alt_conversep) + apply(subst (asm) (1 3 4 5) eq_alt) + apply(subst (asm) (1) eq_alt_conversep2) + apply(unfold prod.rel_conversep rel_gpv''_conversep prod.rel_Grp rel_gpv''_Grp) + apply(force simp add: rel_fun_def Grp_def map_gpv_conv_map_gpv' map_fun_def[abs_def] o_def apfst_def) + done + +lemma interaction_bound_try_gpv: + fixes "consider" defines "ib \ interaction_bound consider" + shows "interaction_bound consider (try_gpv gpv gpv') \ ib gpv + ib gpv'" +proof(induction arbitrary: gpv gpv' rule: interaction_bound_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step interaction_bound') + show ?case unfolding ib_def + apply(clarsimp simp add: generat.map_comp image_image o_def case_map_generat cong del: generat.case_cong split!: if_split generat.split intro!: SUP_least) + subgoal + apply(subst interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply clarsimp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(simp add: iadd_Suc) + apply(subst Sup_image_eadd1[symmetric]) + apply simp + apply(rule SUP_mono) + apply simp + apply(rule exI) + apply(rule step.IH[unfolded ib_def]) + done + subgoal + apply(subst interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply clarsimp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply simp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(rule step.IH[unfolded ib_def]) + done + subgoal + apply(subst interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply clarsimp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(simp add: iadd_Suc) + apply(subst Sup_image_eadd1[symmetric]) + apply simp + apply(rule SUP_mono) + apply simp + apply(rule exI) + apply(rule step.IH[unfolded ib_def]) + done + subgoal + apply(subst interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply clarsimp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(subst Sup_image_eadd1[symmetric]) + apply simp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(rule step.IH[unfolded ib_def]) + done + subgoal + apply(subst (2) interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd2[symmetric]) + apply clarsimp + apply simp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(simp add: iadd_Suc_right) + apply(subst Sup_image_eadd2[symmetric]) + apply clarsimp + apply(rule SUP_mono) + apply clarsimp + apply(rule exI) + apply(rule order_trans) + apply(rule step.hyps) + apply(rule enat_le_plus_same) + done + subgoal + apply(subst (2) interaction_bound.simps) + apply simp + apply(subst Sup_image_eadd2[symmetric]) + apply clarsimp + apply simp + apply(rule SUP_upper2) + apply(rule rev_image_eqI) + apply simp + apply simp + apply(subst Sup_image_eadd2[symmetric]) + apply clarsimp + apply(rule SUP_upper2) + apply(rule imageI) + apply simp + apply(rule order_trans) + apply(rule step.hyps) + apply(rule enat_le_plus_same) + done + done +qed + +lemma interaction_bounded_by_try_gpv [interaction_bound]: + "interaction_bounded_by consider (try_gpv gpv1 gpv2) (bound1 + bound2)" + if "interaction_bounded_by consider gpv1 bound1" "interaction_bounded_by consider gpv2 bound2" + using that interaction_bound_try_gpv[of "consider" gpv1 gpv2] + by(simp add: interaction_bounded_by.simps)(meson add_left_mono_trans add_right_mono le_left_mono) + + +subsection \term \gpv_stop\\ + +lemma interaction_bounded_by_gpv_stop [interaction_bound]: + assumes "interaction_bounded_by consider gpv n" + shows "interaction_bounded_by consider (gpv_stop gpv) n" + using assms by(simp add: interaction_bounded_by.simps) + +context includes \.lifting begin + +lift_definition stop_\ :: "('a, 'b) \ \ ('a, 'b option) \" is + "\resp x. if (resp x = {}) then {} else insert None (Some ` resp x)" . + +lemma outs_stop_\ [simp]: "outs_\ (stop_\ \) = outs_\ \" + by transfer auto + +lemma responses_stop_\ [simp]: + "responses_\ (stop_\ \) x = (if x \ outs_\ \ then insert None (Some ` responses_\ \ x) else {})" + by transfer auto + +lemma stop_\_full [simp]: "stop_\ \_full = \_full" + by transfer(auto simp add: fun_eq_iff notin_range_Some) + +lemma stop_\_uniform [simp]: + "stop_\ (\_uniform A B) = (if B = {} then \ else \_uniform A (insert None (Some ` B)))" + unfolding bot_\_def by transfer(simp add: fun_eq_iff) + +lifting_update \.lifting +lifting_forget \.lifting + +end + +lemma stop_\_bot [simp]: "stop_\ \ = \" + by(simp only: bot_\_def stop_\_uniform)(simp) + +lemma WT_gpv_stop [simp, WT_intro]: "stop_\ \ \g gpv_stop gpv \" if "\ \g gpv \" + using that by(coinduction arbitrary: gpv)(auto 4 3 dest: WT_gpvD) + +lemma expectation_gpv_stop: + fixes fail and gpv :: "('a, 'b, 'c) gpv" + assumes WT: "\ \g gpv \" + and fail: "fail \ c" + shows "expectation_gpv fail (stop_\ \) (\_. c) (gpv_stop gpv) = expectation_gpv fail \ (\_. c) gpv" (is "?lhs = ?rhs") +proof(rule antisym) + show "expectation_gpv fail (stop_\ \) (\_. c) (gpv_stop gpv) \ expectation_gpv fail \ (\_. c) gpv" + using WT + proof(induction arbitrary: gpv rule: parallel_fixp_induct_1_1[OF complete_lattice_partial_function_definitions complete_lattice_partial_function_definitions expectation_gpv.mono expectation_gpv.mono expectation_gpv_def expectation_gpv_def, case_names adm bottom step]) + case adm show ?case by simp + case bottom show ?case by simp + case (step f g) + then show ?case + apply(simp add: pmf_map_spmf_None measure_spmf_return_spmf nn_integral_return) + apply(rule disjI2 nn_integral_mono_AE)+ + apply(auto split!: generat.split simp add: image_image dest: WT_gpvD intro!: le_infI2 INF_mono) + done + qed + + define stop :: "('a option, 'b, 'c option) gpv \ _" where "stop = expectation_gpv fail (stop_\ \) (\_. c)" + show "?rhs \ stop (gpv_stop gpv)" using WT + proof(induction arbitrary: gpv rule: expectation_gpv_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step expectation_gpv') + have "expectation_gpv' gpv' \ c" if "\ \g gpv' \" for gpv' + using expectation_gpv_const_le[of \ gpv' fail c] fail step.hyps(1)[of gpv'] that + by(simp add: max_def split: if_split_asm) + then show ?case using step unfolding stop_def + apply(subst expectation_gpv.simps) + apply(simp add: pmf_map_spmf_None) + apply(rule disjI2 nn_integral_mono_AE)+ + apply(clarsimp split!: generat.split simp add: image_image) + subgoal by(auto 4 3 simp add: in_outs_\_iff_responses_\ dest: WT_gpv_ContD intro: INF_lower2) + subgoal by(auto intro!: INF_mono rev_bexI dest: WT_gpvD) + done + qed +qed + +lemma pgen_lossless_gpv_stop: + fixes fail and gpv :: "('a, 'b, 'c) gpv" + assumes WT: "\ \g gpv \" + and fail: "fail \ 1" + shows "pgen_lossless_gpv fail (stop_\ \) (gpv_stop gpv) = pgen_lossless_gpv fail \ gpv" + by(simp add: pgen_lossless_gpv_def expectation_gpv_stop assms) + +lemma pfinite_gpv_stop [simp]: + "pfinite_gpv (stop_\ \) (gpv_stop gpv) \ pfinite_gpv \ gpv" if "\ \g gpv \" + using that by(simp add: pgen_lossless_gpv_stop) + +lemma plossless_gpv_stop [simp]: + "plossless_gpv (stop_\ \) (gpv_stop gpv) \ plossless_gpv \ gpv" if "\ \g gpv \" + using that by(simp add: pgen_lossless_gpv_stop) + +lemma results_gpv_stop_SomeD: "Some x \ results_gpv (stop_\ \) (gpv_stop gpv) \ x \ results_gpv \ gpv" + by(induction gpv'\"gpv_stop gpv" arbitrary: gpv rule: results_gpv.induct) + (auto 4 3 intro: results_gpv.intros split: if_split_asm) + +lemma Some_in_results'_gpv_gpv_stopD: "Some xy \ results'_gpv (gpv_stop gpv) \ xy \ results'_gpv gpv" + unfolding results_gpv_\_full[symmetric] + by(rule results_gpv_stop_SomeD) simp + +subsection \term \exception_\\\ + +datatype 's exception = Fault | OK (ok: 's) + +lemma inj_on_OK [simp]: "inj_on OK A" + by(auto simp add: inj_on_def) + +function join_exception :: "'a exception \ 'b exception \ ('a \ 'b) exception" where + "join_exception Fault _ = Fault" +| "join_exception _ Fault = Fault" +| "join_exception (OK a) (OK b) = OK (a, b)" + by pat_completeness auto +termination by lexicographic_order + +primrec merge_exception :: "'a exception + 'b exception \ ('a + 'b) exception" where + "merge_exception (Inl x) = map_exception Inl x" +| "merge_exception (Inr y) = map_exception Inr y" + +fun option_of_exception :: "'a exception \ 'a option" where + "option_of_exception Fault = None" +| "option_of_exception (OK x) = Some x" + +fun exception_of_option :: "'a option \ 'a exception" where + "exception_of_option None = Fault" +| "exception_of_option (Some x) = OK x" + + +lemma option_of_exception_exception_of_option [simp]: "option_of_exception (exception_of_option x) = x" + by(cases x) simp_all + +lemma exception_of_option_option_of_exception [simp]: "exception_of_option (option_of_exception x) = x" + by(cases x) simp_all + +lemma case_exception_of_option [simp]: "case_exception f g (exception_of_option x) = case_option f g x" + by(simp split: exception.split option.split) + +lemma case_option_of_exception [simp]: "case_option f g (option_of_exception x) = case_exception f g x" + by(simp split: exception.split option.split) + +lemma surj_exception_of_option [simp]: "surj exception_of_option" + by(rule surjI[where f="option_of_exception"])(simp) + +lemma surj_option_of_exception [simp]: "surj option_of_exception" + by(rule surjI[where f="exception_of_option"])(simp) + +lemma case_map_exception [simp]: "case_exception f g (map_exception h x) = case_exception f (g \ h) x" + by(simp split: exception.split) + +definition exception_\ :: "('a, 'b) \ \ ('a, 'b exception) \" where + "exception_\ \ = map_\ id exception_of_option (stop_\ \)" + +lemma outs_exception_\ [simp]: "outs_\ (exception_\ \) = outs_\ \" + by(simp add: exception_\_def) + +lemma responses_exception_\ [simp]: + "responses_\ (exception_\ \) x = (if x \ outs_\ \ then insert Fault (OK ` responses_\ \ x) else {})" + by(simp add: exception_\_def image_image) + +lemma map_\_full [simp]: "map_\ f g \_full = \_uniform UNIV (range g)" + unfolding \_uniform_UNIV[symmetric] map_\_\_uniform by simp + +lemma exception_\_full [simp]: "exception_\ \_full = \_full" + unfolding exception_\_def by simp + +lemma exception_\_uniform [simp]: + "exception_\ (\_uniform A B) = (if B = {} then \ else \_uniform A (insert Fault (OK ` B)))" + by(simp add: exception_\_def image_image) + +lemma option_of_exception_\ [simp]: "map_\ id option_of_exception (exception_\ \) = stop_\ \" + by(simp add: exception_\_def o_def id_def[symmetric]) + +lemma exception_of_option_\ [simp]: "map_\ id exception_of_option (stop_\ \) = exception_\ \" + by(simp add: exception_\_def) + +subsection \inline\ + +context raw_converter_invariant begin + +context + fixes gpv :: "('a, 'call, 'ret) gpv" + assumes gpv: "plossless_gpv \ gpv" "\ \g gpv \" +begin + +lemma lossless_spmf_inline1: + assumes lossless: "\s x. \ x \ outs_\ \; I s \ \ lossless_spmf (the_gpv (callee s x))" + and I: "I s" + shows "lossless_spmf (inline1 callee gpv s)" +proof - + have "1 = expectation_gpv 0 \ (\_. 1) gpv" using gpv by(simp add: pgen_lossless_gpv_def) + also have "\ \ weight_spmf (inline1 callee gpv s)" using gpv(2) I + proof(induction arbitrary: gpv s rule: expectation_gpv_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step expectation_gpv') + { fix out c + assume IO: "IO out c \ set_spmf (the_gpv gpv)" + with step.prems have out: "out \ outs_\ \" by(auto dest: WT_gpvD) + from out[unfolded in_outs_\_iff_responses_\] obtain input where input: "input \ responses_\ \ out" by auto + from out have "(\r\responses_\ \ out. expectation_gpv' (c r)) = \\<^sup>+ x. (\r\responses_\ \ out. expectation_gpv' (c r)) \measure_spmf (the_gpv (callee s out))" + using lossless \I s\ by(simp add: lossless_spmf_def measure_spmf.emeasure_eq_measure) + also have "\ \ \\<^sup>+ generat. (case generat of Pure (r, s') \ weight_spmf (inline1 callee (c r) s') | _ \ 1) \measure_spmf (the_gpv (callee s out))" + apply(intro nn_integral_mono_AE) + apply(clarsimp split!: generat.split) + subgoal Pure + apply(rule INF_lower2) + apply(fastforce dest: results_callee[OF out \I s\, THEN subsetD, OF results_gpv.Pure]) + apply(rule step.IH) + apply(fastforce intro: WT_gpvD[OF step.prems(1) IO] dest: results_callee[OF out \I s\, THEN subsetD, OF results_gpv.Pure]) + apply(fastforce dest: results_callee[OF out \I s\, THEN subsetD, OF results_gpv.Pure]) + done + subgoal IO + apply(rule INF_lower2[OF input]) + apply(rule order_trans) + apply(rule step.hyps) + apply(rule order_trans) + apply(rule expectation_gpv_const_le) + apply(rule WT_gpvD[OF step.prems(1) IO]) + apply(simp_all add: input) + done + done + finally have "(\r\responses_\ \ out. expectation_gpv' (c r)) \ \" . } + then show ?case using step.prems + apply(subst inline1.simps) + apply(simp add: measure_spmf.emeasure_eq_measure[symmetric]) + apply(simp add: measure_spmf_bind) + apply(subst emeasure_bind[where N="count_space UNIV"]) + apply(simp add: space_measure_spmf) + apply(simp add: o_def) + apply(simp) + apply(rule nn_integral_mono_AE) + apply(clarsimp split!: generat.split simp add: measure_spmf_return_spmf space_measure_spmf) + apply(simp add: measure_spmf_bind) + apply(subst emeasure_bind[where N="count_space UNIV"]) + apply(simp add: space_measure_spmf) + apply(simp add: o_def) + apply(simp) + apply (simp add: measure_spmf.emeasure_eq_measure) + apply(subst generat.case_distrib[where h="\x. measure (measure_spmf x) _"]) + apply(simp add: split_def measure_spmf_return_spmf space_measure_spmf measure_return cong del: generat.case_cong) + done + qed + finally show ?thesis using weight_spmf_le_1[of "inline1 callee gpv s"] by(simp add: lossless_spmf_def) +qed + +end + +end + +lemma (in raw_converter_invariant) inline1_try_gpv: + defines "inline1' \ inline1" + assumes WT: "\ \g gpv \" + and pfinite: "pfinite_gpv \ gpv" + and f: "\s. I s \ f (x, s) = z" + and lossless: "\s x. \ x \ outs_\ \; I s \ \ colossless_gpv \' (callee s x)" + and I: "I s" + shows "map_spmf (map_sum f id) (inline1 callee (try_gpv gpv (Done x)) s) = + try_spmf (map_spmf (map_sum f (\(out, c, rpv). (out, c, \input. try_gpv (rpv input) (Done x)))) (inline1' callee gpv s)) (return_spmf (Inl z))" + (is "?lhs = ?rhs") +proof - + have le: "ord_spmf (=) ?lhs ?rhs" using WT I + proof(induction arbitrary: gpv s rule: inline1_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step inline1'') + show ?case using step.prems unfolding inline1'_def + apply(subst inline1.simps) + apply(simp add: bind_map_spmf map_bind_spmf o_def) + apply(simp add: try_spmf_def) + apply(subst bind_spmf_pmf_assoc) + apply(simp add: bind_map_pmf) + apply(subst (3) bind_spmf_def) + apply(simp add: bind_assoc_pmf) + apply(rule rel_pmf_bindI[where R="eq_onp (\x. x \ set_pmf (the_gpv gpv))"]) + apply(rule pmf.rel_refl_strong) + apply(simp add: eq_onp_def) + apply(clarsimp simp add: eq_onp_def bind_return_pmf f split!: option.split generat.split) + subgoal for out c + apply(simp add: in_set_spmf[symmetric] bind_map_pmf map_bind_spmf) + apply(subst option.case_distrib[where h=return_pmf, symmetric, abs_def]) + apply(fold map_pmf_def) + apply(simp add: bind_spmf_def map_bind_pmf) + apply(rule rel_pmf_bindI[where R="eq_onp (\x. x \ set_pmf (the_gpv (callee s out)))"]) + apply(rule pmf.rel_refl_strong) + apply(simp add: eq_onp_def) + apply(simp add: in_set_spmf[symmetric] bind_map_pmf map_bind_spmf eq_onp_def split!: option.split generat.split) + apply(rule spmf.leq_trans) + apply(rule step.IH[unfolded inline1'_def]) + subgoal + by(auto dest: results_callee[THEN subsetD, OF _ _ results_gpv.Pure, rotated -1] WT_gpvD) + subgoal + by(auto dest: results_callee[THEN subsetD, OF _ _ results_gpv.Pure, rotated -1] WT_gpvD) + apply(simp add: try_spmf_def) + apply(subst option.case_distrib[where h=return_pmf, symmetric, abs_def]) + apply(fold map_pmf_def) + apply simp + done + done + qed + have "lossless_spmf ?lhs" using I + apply simp + apply(rule lossless_spmf_inline1) + apply(rule plossless_gpv_try_gpvI) + apply(rule pfinite) + apply simp + apply(rule WT_gpv_try_gpvI) + apply(rule WT) + apply simp + apply(rule colossless_gpv_lossless_spmfD[OF lossless]) + apply simp_all + done + from ord_spmf_lossless_spmfD1[OF le this] show ?thesis by(simp add: spmf_rel_eq) +qed + +lemma (in raw_converter_invariant) inline_try_gpv: + assumes WT: "\ \g gpv \" + and pfinite: "pfinite_gpv \ gpv" + and f: "\s. I s \ f (x, s) = z" + and lossless: "\s x. \ x \ outs_\ \; I s \ \ colossless_gpv \' (callee s x)" + and I: "I s" + shows "eq_\_gpv (=) \' (map_gpv f id (inline callee (try_gpv gpv (Done x)) s)) (try_gpv (map_gpv f id (inline callee gpv s)) (Done z))" + (is "eq_\_gpv _ _ ?lhs ?rhs") + using WT pfinite I +proof(coinduction arbitrary: gpv s rule: eq_\_gpv_coinduct_bind) + case (eq_\_gpv gpv s) + show "?case TYPE(('ret \ 's) option) TYPE(('ret \ 's) option)" (is "rel_spmf (eq_\_generat _ _ ?X) ?lhs ?rhs") + proof - + have "?lhs = map_spmf + (\x. case x of Inl rs \ Pure rs | Inr (out, oracle, rpv) \ IO out (\input. + map_gpv f id (bind_gpv (try_gpv (map_gpv Some id (oracle input)) (Done None)) (\xy. case xy of None \ Fail | Some (x, y) \ inline callee (rpv x) y)))) + (map_spmf (map_sum f id) (inline1 callee (TRY gpv ELSE Done x) s))" + (is "_ = map_spmf ?f ?lhs2") + by(auto simp add: gpv.map_sel inline_sel spmf.map_comp o_def bind_gpv_try_gpv_map_Some intro!: map_spmf_cong[OF refl] split: sum.split) + also from eq_\_gpv + have "?lhs2 = TRY map_spmf (map_sum f (\(out, c, rpv). (out, c, \input. TRY rpv input ELSE Done x))) (inline1 callee gpv s) ELSE return_spmf (Inl z)" + by(intro inline1_try_gpv)(auto intro: f lossless) + also have "\ = map_spmf (\y. case y of None \ Inl z | Some x' \ map_sum f (\(out, c, rpv). (out, c, \input. try_gpv (rpv input) (Done x))) x') + (try_spmf (map_spmf Some (inline1 callee gpv s)) (return_spmf None))" + (is "_ = ?lhs3") by(simp add: map_try_spmf spmf.map_comp o_def) + also have "?rhs = map_spmf (\y. case y of None \ Pure z | Some (Inl x) \ Pure (f x) + | Some (Inr (out, oracle, rpv)) \ IO out (\input. try_gpv (map_gpv f id (bind_gpv (oracle input) (\(x, y). inline callee (rpv x) y))) (Done z))) + (try_spmf (map_spmf Some (inline1 callee gpv s)) (return_spmf None))" + by(auto simp add: gpv.map_sel inline_sel spmf.map_comp o_def generat.map_comp spmf_rel_map map_try_spmf intro!: try_spmf_cong map_spmf_cong split: sum.split) + moreover have "rel_spmf (eq_\_generat (=) \' ?X) (map_spmf ?f ?lhs3) \" + apply(clarsimp simp add: gpv.map_sel inline_sel spmf.map_comp o_def generat.map_comp spmf_rel_map intro!: rel_spmf_reflI) + apply(erule disjE) + subgoal + apply(clarsimp split!: generat.split sum.split simp add: map_gpv_id_bind_gpv) + apply(subst (3) try_gpv_bind_gpv) + apply(rule conjI) + apply(erule WT_gpv_inline1[OF _ eq_\_gpv(1,3)]) + apply(rule strip)+ + apply(rule disjI2)+ + subgoal for out rpv rpv' input + apply(rule exI) + apply(rule exI) + apply(rule exI[where x="\x y. x = y \ y \ results_gpv \' (TRY map_gpv Some id (rpv input) ELSE Done None)"]) + apply(rule exI conjI refl)+ + apply(rule eq_\_gpv_reflI) + apply(simp add: eq_onp_def) + apply(rule WT_intro) + apply simp + apply(erule (1) WT_gpv_inline1[OF _ eq_\_gpv(1,3)]) + apply simp + apply(rule rel_funI) + apply(clarsimp simp add: eq_onp_def split: if_split_asm) + subgoal + apply(rule exI conjI refl)+ + apply(drule (2) WT_gpv_inline1(3)[OF _ eq_\_gpv(1,3)]) + apply simp + apply(frule (2) WT_gpv_inline1(3)[OF _ eq_\_gpv(1,3)]) + apply(drule (2) inline1_in_sub_gpvs[OF _ _ _ eq_\_gpv(1,3)]) + apply clarsimp + apply(erule pfinite_gpv_sub_gpvs[OF eq_\_gpv(2) _ eq_\_gpv(1)]) + done + subgoal + apply(erule disjE; clarsimp) + apply(rule exI conjI refl)+ + apply(drule (2) WT_gpv_inline1(3)[OF _ eq_\_gpv(1,3)]) + apply simp + apply(frule (2) WT_gpv_inline1(3)[OF _ eq_\_gpv(1,3)]) + apply(drule (2) inline1_in_sub_gpvs[OF _ _ _ eq_\_gpv(1,3)]) + apply clarsimp + apply(erule pfinite_gpv_sub_gpvs[OF eq_\_gpv(2) _ eq_\_gpv(1)]) + apply(erule notE) + apply(drule inline1_in_sub_gpvs_callee[OF _ eq_\_gpv(1,3)]) + apply clarify + apply(drule (1) bspec) + apply(erule colossless_gpv_sub_gpvs[rotated]) + apply(rule lossless; simp) + done + done + done + subgoal by(clarsimp split: if_split_asm) + done + ultimately show ?thesis by(simp only:) + qed +qed + + +definition cr_prod2 :: "'a \ ('b \ 'c \ bool) \ 'b \ 'a \ 'c \ bool" where + "cr_prod2 x A = (\b (a, c). A b c \ x = a)" + +lemma cr_prod2_simps [simp]: "cr_prod2 x A a (b, c) \ A a c \ x = b" +by(simp add: cr_prod2_def) + +lemma cr_prod2I: "A a b \ cr_prod2 x A a (x, b)" by simp + +lemma cr_prod2_Grp: "cr_prod2 x (BNF_Def.Grp A f) = BNF_Def.Grp A (\b. (x, f b))" +by(auto simp add: Grp_def fun_eq_iff) + +(* A stronger version of the existing extend_state_oracle_transfer *) +lemma extend_state_oracle_transfer': includes lifting_syntax shows + "((S ===> C ===> rel_spmf (rel_prod R S)) ===> cr_prod2 s S ===> C ===> rel_spmf (rel_prod R (cr_prod2 s S))) (\oracle. oracle) extend_state_oracle" +unfolding extend_state_oracle_def[abs_def] +apply(rule rel_funI)+ +apply clarsimp +apply(drule (1) rel_funD)+ +apply(auto simp add: spmf_rel_map split_def dest: rel_funD intro: rel_spmf_mono) +done + +(* @Reza: This is a proof based on parametricity, + using the relation cr_prod2 s (=) to fix the first component of the extended callee state to s *) + +lemma exec_gpv_extend_state_oracle: + "exec_gpv (extend_state_oracle callee) gpv (s, s') = + map_spmf (\(x, s''). (x, (s, s''))) (exec_gpv callee gpv s')" + using exec_gpv_parametric'[THEN rel_funD, OF extend_state_oracle_transfer'[THEN rel_funD], of "(=)" "(=)" "(=)" callee callee "(=)" s] + unfolding relator_eq rel_gpv''_eq + apply(clarsimp simp add: rel_fun_def) + apply(unfold eq_alt cr_prod2_Grp prod.rel_Grp option.rel_Grp pmf.rel_Grp) + apply(simp add: Grp_def map_prod_def) + apply(blast intro: sym) + done + + + + + +section \Material for Constructive Crypto\ + +lemma WT_resource_\_uniform_UNIV [simp]: "\_uniform A UNIV \res res \" + by(coinduction arbitrary: res) auto + +lemma WT_converter_of_callee_invar: + assumes WT: "\s q. \ q \ outs_\ \; I s \ \ \' \g callee s q \" + and res: "\s q r s'. \ (r, s') \ results_gpv \' (callee s q); q \ outs_\ \; I s \ \ r \ responses_\ \ q \ I s'" + and I: "I s" + shows "\, \' \\<^sub>C converter_of_callee callee s \" + using I by(coinduction arbitrary: s)(auto simp add: WT res) + +lemma eq_\_gpv_eq_OO: + assumes "eq_\_gpv (=) \ gpv gpv'" "eq_\_gpv A \ gpv' gpv''" + shows "eq_\_gpv A \ gpv gpv''" + using eq_\_gpv_relcompp[THEN fun_cong, THEN fun_cong, THEN iffD2, OF relcomppI, OF assms] + by(simp add: eq_OO) + +lemma eq_\_gpv_eq_OO2: + assumes "eq_\_gpv (=) \ gpv'' gpv'" "eq_\_gpv A \ gpv gpv'" + shows "eq_\_gpv A \ gpv gpv''" + using eq_\_gpv_relcompp[where A'="conversep (=)", THEN fun_cong, THEN fun_cong, THEN iffD2, OF relcomppI, OF assms(2)] assms(1) + unfolding eq_\_gpv_conversep by(simp add: OO_eq) + +lemma eq_\_gpv_try_gpv_cong: + assumes "eq_\_gpv A \ gpv1 gpv1'" + and "eq_\_gpv A \ gpv2 gpv2'" + shows "eq_\_gpv A \ (try_gpv gpv1 gpv2) (try_gpv gpv1' gpv2')" + using assms(1) + apply(coinduction arbitrary: gpv1 gpv1') + using assms(2) + apply(fastforce simp add: spmf_rel_map intro!: rel_spmf_try_spmf dest: eq_\_gpvD elim!: rel_spmf_mono_strong eq_\_generat.cases) + done + +lemma eq_\_gpv_map_gpv': + assumes "eq_\_gpv (BNF_Def.vimage2p f f' A) (map_\ g h \) gpv1 gpv2" + shows "eq_\_gpv A \ (map_gpv' f g h gpv1) (map_gpv' f' g h gpv2)" + using assms +proof(coinduction arbitrary: gpv1 gpv2) + case eq_\_gpv + from this[THEN eq_\_gpvD] show ?case + apply(simp add: spmf_rel_map) + apply(erule rel_spmf_mono) + apply(auto 4 4 simp add: BNF_Def.vimage2p_def elim!: eq_\_generat.cases) + done +qed + +lemma eq_\_converter_map_converter: + assumes "map_\ (inv_into UNIV f) (inv_into UNIV g) \, map_\ f' g' \' \\<^sub>C conv1 \ conv2" + and "inj f" "surj g" + shows "\, \' \\<^sub>C map_converter f g f' g' conv1 \ map_converter f g f' g' conv2" + using assms(1) +proof(coinduction arbitrary: conv1 conv2) + case eq_\_converter + from this(2) have "f q \ outs_\ (map_\ (inv_into UNIV f) (inv_into UNIV g) \)" using assms(2) by simp + from eq_\_converter(1)[THEN eq_\_converterD, OF this] show ?case using assms(2,3) + apply simp + apply(rule eq_\_gpv_map_gpv') + apply(simp add: BNF_Def.vimage2p_def prod.rel_map) + apply(erule eq_\_gpv_mono') + apply(auto 4 4 simp add: eq_onp_def surj_f_inv_f) + done +qed + +lemma resource_of_oracle_run_resource: "resource_of_oracle run_resource res = res" + by(coinduction arbitrary: res)(auto simp add: rel_fun_def spmf_rel_map intro!: rel_spmf_reflI) + +lemma connect_map_gpv': + "connect (map_gpv' f g h adv) res = map_spmf f (connect adv (map_resource g h res))" + unfolding connect_def + by(subst (3) resource_of_oracle_run_resource[symmetric]) + (simp add: exec_gpv_map_gpv' map_resource_resource_of_oracle spmf.map_comp exec_gpv_resource_of_oracle) + +primcorec fail_resource :: "('a, 'b) resource" where + "run_resource fail_resource = (\_. return_pmf None)" + +lemma WT_fail_resource [WT_intro]: "\ \res fail_resource \" + by(rule WT_resourceI) simp + +context fixes y :: "'b" begin + +primcorec const_resource :: "('a, 'b) resource" where + "run_resource const_resource = (\_. map_spmf (map_prod id (\_. const_resource)) (return_spmf (y, ())))" + +end + +lemma const_resource_sel [simp]: "run_resource (const_resource y) = (\_. return_spmf (y, const_resource y))" + by simp + +declare const_resource.sel [simp del] + +lemma lossless_const_resource [simp]: "lossless_resource \ (const_resource y)" + by(coinduction) simp + +lemma WT_const_resource [simp]: + "\ \res const_resource y \ \ (\x\outs_\ \. y \ responses_\ \ x)" (is "?lhs \ ?rhs") +proof(intro iffI ballI) + show "y \ responses_\ \ x" if ?lhs and "x \ outs_\ \" for x using WT_resourceD[OF that] by auto + show ?lhs if ?rhs using that by(coinduction)(auto) +qed + +context fixes y :: "'b" begin + +primcorec const_converter :: "('a, 'b, 'c, 'd) converter" where + "run_converter const_converter = (\_. map_gpv (map_prod id (\_. const_converter)) id (Done (y, ())))" + +end + +lemma const_converter_sel [simp]: "run_converter (const_converter y) = (\_. Done (y, const_converter y))" + by simp + +lemma attach_const_converter [simp]: "attach (const_converter y) res = const_resource y" + by(coinduction)(simp add: rel_fun_def) + +declare const_converter.sel [simp del] + +lemma comp_const_converter [simp]: "comp_converter (const_converter x) conv = const_converter x" + by(coinduction)(simp add: rel_fun_def) + +lemma interaction_bounded_const_converter [simp, interaction_bound]: + "interaction_any_bounded_converter (const_converter Fault) bound" + by(coinduction) simp + + +primcorec merge_exception_converter :: "('a, ('b + 'c) exception, 'a, 'b exception + 'c exception) converter" where + "run_converter merge_exception_converter = + (\x. map_gpv (map_prod id (\conv. case conv of None \ merge_exception_converter | Some conv' \ conv')) id ( + Pause x (\y. Done (case merge_exception y of Fault \ (Fault, Some (const_converter Fault)) + | OK y' \ (OK y', None)))))" + +lemma merge_exception_converter_sel [simp]: + "run_converter merge_exception_converter x = + Pause x (\y. Done (case merge_exception y of Fault \ (Fault, const_converter Fault) | OK y' \ (OK y', merge_exception_converter)))" + by(simp add: o_def fun_eq_iff split: exception.split) + +declare merge_exception_converter.sel[simp del] + +lemma plossless_const_converter[simp]: "plossless_converter \ \' (const_converter x)" + by(coinduction) auto + +lemma plossless_merge_exception_converter [simp]: + "plossless_converter (exception_\ (\ \\<^sub>\ \')) (exception_\ \ \\<^sub>\ exception_\ \') merge_exception_converter" + by(coinduction) auto + +lemma WT_const_converter [WT_intro, simp]: + "\, \' \\<^sub>C const_converter x \" if "\q \ outs_\ \. x \ responses_\ \ q" + by(coinduction)(auto simp add: that) + +lemma WT_merge_exception_converter [WT_intro, simp]: + "exception_\ (\1' \\<^sub>\ \2'), exception_\ \1' \\<^sub>\ exception_\ \2' \\<^sub>C merge_exception_converter \" + by(coinduction) auto + +lemma inline_left_gpv_merge_exception_converter: + "bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (left_gpv gpv))) merge_exception_converter) (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv')) = + bind_gpv (left_gpv (map_gpv' id id option_of_exception (gpv_stop gpv))) (\x. case x of None \ Fail | Some x' \ Done (x, merge_exception_converter))" + apply(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + apply(simp add: bind_gpv.sel inline_sel map_bind_spmf bind_map_spmf del: bind_gpv_sel') + apply(subst inline1_unfold) + apply(clarsimp simp add: bind_map_spmf intro!: rel_spmf_bind_reflI simp add: generat.map_comp case_map_generat o_def split!: generat.split intro!: rel_funI) + subgoal for gpv out c input by(cases input; auto split!: exception.split) + done + +lemma inline_right_gpv_merge_exception_converter: + "bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (right_gpv gpv))) merge_exception_converter) (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv')) = + bind_gpv (right_gpv (map_gpv' id id option_of_exception (gpv_stop gpv))) (\x. case x of None \ Fail | Some x' \ Done (x, merge_exception_converter))" + apply(coinduction arbitrary: gpv rule: gpv.coinduct_strong) + apply(simp add: bind_gpv.sel inline_sel map_bind_spmf bind_map_spmf del: bind_gpv_sel') + apply(subst inline1_unfold) + apply(clarsimp simp add: bind_map_spmf intro!: rel_spmf_bind_reflI simp add: generat.map_comp case_map_generat o_def split!: generat.split intro!: rel_funI) + subgoal for gpv out c input by(cases input; auto split!: exception.split) + done + +subsection \@{theory Constructive_Cryptography.Wiring}\ + +abbreviation (input) + id_wiring :: "('a, 'b, 'a, 'b) wiring" ("1\<^sub>w") + where + "id_wiring \ (id, id)" + +definition + swap_lassocr\<^sub>w :: "('a + 'b + 'c, 'd + 'e + 'f, 'b + 'a + 'c, 'e + 'd + 'f) wiring" + where + "swap_lassocr\<^sub>w \ rassocl\<^sub>w \\<^sub>w ((swap\<^sub>w |\<^sub>w 1\<^sub>w) \\<^sub>w lassocr\<^sub>w)" + +schematic_goal + wiring_swap_lassocr[wiring_intro]: "wiring ?\1 ?\2 swap_lassocr swap_lassocr\<^sub>w" + unfolding swap_lassocr_def swap_lassocr\<^sub>w_def + by(rule wiring_intro)+ + +definition + parallel_wiring\<^sub>w :: "(('a + 'b) + ('c + 'd), ('e + 'f) + ('g + 'h), + ('a + 'c) + ('b + 'd), ('e + 'g) + ('f + 'h)) wiring" + where + "parallel_wiring\<^sub>w \ lassocr\<^sub>w \\<^sub>w ((1\<^sub>w |\<^sub>w swap_lassocr\<^sub>w) \\<^sub>w rassocl\<^sub>w)" + +schematic_goal + wiring_parallel_wiring[wiring_intro]: "wiring ?\1 ?\2 parallel_wiring parallel_wiring\<^sub>w" + unfolding parallel_wiring_def parallel_wiring\<^sub>w_def + by(rule wiring_intro)+ + + +lemma lassocr_inverse: "rassocl\<^sub>C \ lassocr\<^sub>C = 1\<^sub>C" + unfolding rassocl\<^sub>C_def lassocr\<^sub>C_def + apply(simp add: comp_converter_map1_out comp_converter_map_converter2 comp_converter_id_right) + apply(subst map_converter_id_move_right) + apply(simp add: o_def id_def[symmetric]) + done + +lemma rassocl_inverse: "lassocr\<^sub>C \ rassocl\<^sub>C = 1\<^sub>C" + unfolding rassocl\<^sub>C_def lassocr\<^sub>C_def + apply(simp add: comp_converter_map1_out comp_converter_map_converter2 comp_converter_id_right) + apply(subst map_converter_id_move_right) + apply(simp add: o_def id_def[symmetric]) + done + +lemma swap_sum_swap_sum [simp]: "swap_sum (swap_sum x) = x" + by(cases x) simp_all + +lemma inj_on_lsumr [simp]: "inj_on lsumr A" + by(auto simp add: inj_on_def elim: lsumr.elims) + +lemma inj_on_rsuml [simp]: "inj_on rsuml A" + by(auto simp add: inj_on_def elim: rsuml.elims) + +lemma bij_lsumr [simp]: "bij lsumr" + by(rule o_bij[where g=rsuml]) auto + +lemma bij_swap_sum [simp]: "bij swap_sum" + by(rule o_bij[where g=swap_sum]) auto + +lemma bij_rsuml [simp]: "bij rsuml" + by(rule o_bij[where g=lsumr]) auto + +lemma bij_lassocr_swap_sum [simp]: "bij lassocr_swap_sum" + unfolding lassocr_swap_sum_def + by(simp add: bij_comp) + +lemma inj_lassocr_swap_sum [simp]: "inj lassocr_swap_sum" + by(simp add: bij_is_inj) + +lemma inv_rsuml [simp]: "inv_into UNIV rsuml = lsumr" + by(rule inj_imp_inv_eq) auto + +lemma inv_lsumr [simp]: "inv_into UNIV lsumr = rsuml" + by(rule inj_imp_inv_eq) auto + +lemma lassocr_swap_sum_inverse [simp]: "lassocr_swap_sum (lassocr_swap_sum x) = x" + by(simp add: lassocr_swap_sum_def sum.map_comp o_def id_def[symmetric] sum.map_id) + +lemma inv_lassocr_swap_sum [simp]: "inv_into UNIV lassocr_swap_sum = lassocr_swap_sum" + by(rule inj_imp_inv_eq)(simp_all add: sum.map_comp sum.inj_map bij_def surj_iff sum.map_id) + +lemma swap_inverse: "swap\<^sub>C \ swap\<^sub>C = 1\<^sub>C" + unfolding swap\<^sub>C_def + apply(simp add: comp_converter_map1_out comp_converter_map_converter2 comp_converter_id_right) + apply(subst map_converter_id_move_right) + apply(simp add: o_def id_def[symmetric]) + done + +lemma swap_lassocr_inverse: "\1 \\<^sub>\ (\2 \\<^sub>\ \3), \1 \\<^sub>\ (\2 \\<^sub>\ \3) \\<^sub>C swap_lassocr \ swap_lassocr \ 1\<^sub>C" + (is "?\,_ \\<^sub>C ?lhs \ _") +proof - + have "?lhs = (rassocl\<^sub>C \ (swap\<^sub>C |\<^sub>= 1\<^sub>C)) \ (lassocr\<^sub>C \ rassocl\<^sub>C) \ ((swap\<^sub>C |\<^sub>= 1\<^sub>C) \ lassocr\<^sub>C)" + by(simp add: swap_lassocr_def comp_converter_assoc) + also have "\ = rassocl\<^sub>C \ ((swap\<^sub>C \ swap\<^sub>C) |\<^sub>= 1\<^sub>C) \ lassocr\<^sub>C" + unfolding rassocl_inverse comp_converter_id_left + by(simp add: parallel_converter2_comp1_out comp_converter_assoc) + also have "?\,?\ \\<^sub>C \ \ rassocl\<^sub>C \ 1\<^sub>C \ lassocr\<^sub>C" unfolding swap_inverse + by(rule eq_\_converter_reflI eq_\_comp_cong WT_intro parallel_converter2_id_id)+ + also have "rassocl\<^sub>C \ 1\<^sub>C \ lassocr\<^sub>C = 1\<^sub>C" by(simp add: comp_converter_id_left lassocr_inverse) + finally show ?thesis . +qed + +lemma parallel_wiring_inverse: + "(\1 \\<^sub>\ \2) \\<^sub>\ (\3 \\<^sub>\ \4),(\1 \\<^sub>\ \2) \\<^sub>\ (\3 \\<^sub>\ \4) \\<^sub>C parallel_wiring \ parallel_wiring \ 1\<^sub>C" + (is "?\, _ \\<^sub>C ?lhs \ _") +proof - + have "?lhs = (lassocr\<^sub>C \ (1\<^sub>C |\<^sub>= swap_lassocr)) \ (rassocl\<^sub>C \ lassocr\<^sub>C) \ ((1\<^sub>C |\<^sub>= swap_lassocr) \ rassocl\<^sub>C)" + by(simp add: parallel_wiring_def comp_converter_assoc) + also have "\ = (lassocr\<^sub>C \ (1\<^sub>C |\<^sub>= swap_lassocr)) \ (1\<^sub>C |\<^sub>= swap_lassocr) \ rassocl\<^sub>C" + by(simp add: lassocr_inverse comp_converter_id_left) + also have "\ = lassocr\<^sub>C \ (1\<^sub>C |\<^sub>= (swap_lassocr \ swap_lassocr)) \ rassocl\<^sub>C" + by(simp add: parallel_converter2_comp2_out comp_converter_assoc) + also have "?\,?\ \\<^sub>C \ \ lassocr\<^sub>C \ (1\<^sub>C |\<^sub>= 1\<^sub>C) \ rassocl\<^sub>C" + by(rule eq_\_converter_reflI eq_\_comp_cong parallel_converter2_eq_\_cong WT_intro swap_lassocr_inverse)+ + also have "?\,?\ \\<^sub>C lassocr\<^sub>C \ (1\<^sub>C |\<^sub>= 1\<^sub>C) \ rassocl\<^sub>C \ lassocr\<^sub>C \ 1\<^sub>C \ rassocl\<^sub>C" + by(rule eq_\_converter_reflI eq_\_comp_cong parallel_converter2_id_id WT_intro)+ + also have "lassocr\<^sub>C \ 1\<^sub>C \ rassocl\<^sub>C = 1\<^sub>C" by(simp add: comp_converter_id_left rassocl_inverse) + finally show ?thesis . +qed + +\ \ Analogous to @{term \attach_wiring\} in Wiring.thy\ +definition + attach_wiring_right :: " + ('a, 'b, 'c, 'd) wiring \ + ('s \ 'e \ ('f \ 's, 'a, 'b) gpv) \ ('s \ 'e \ ('f \ 's, 'c, 'd) gpv)" + where + "attach_wiring_right = (\(f, g). map_fun id (map_fun id (map_gpv' id f g)))" + +lemma + attach_wiring_right_simps: + "attach_wiring_right (f, g) = map_fun id (map_fun id (map_gpv' id f g))" + by(simp add: attach_wiring_right_def) + +lemma + comp_converter_of_callee_wiring: + assumes wiring: "wiring \2 \3 conv w" + and WT: "\1, \2 \\<^sub>C CNV callee s \" + shows "\1, \3 \\<^sub>C CNV callee s \ conv \ CNV (attach_wiring_right w callee) s" + using wiring +proof cases + case (wiring f g) + from _ wiring(2) have "\1,\3 \\<^sub>C CNV callee s \ conv \ CNV callee s \ map_converter id id f g 1\<^sub>C" + by(rule eq_\_comp_cong)(rule eq_\_converter_reflI[OF WT]) + also have "CNV callee s \ map_converter id id f g 1\<^sub>C = map_converter id id f g (CNV callee s)" + by(subst comp_converter_map_converter2)(simp add: comp_converter_id_right) + also have "\ = CNV (attach_wiring_right w callee) s" + by(simp add: map_converter_of_callee attach_wiring_right_simps wiring(1) prod.map_id0) + finally show ?thesis . +qed + +lemma attach_wiring_right_comp_wiring: + "attach_wiring_right (w1 \\<^sub>w w2) callee = attach_wiring_right w2 (attach_wiring_right w1 callee)" + by(simp add: attach_wiring_right_def comp_wiring_def split_def map_fun_def o_def map_gpv'_comp id_def fun_eq_iff) + +lemma attach_wiring_comp_wiring: + "attach_wiring (w1 \\<^sub>w w2) callee = attach_wiring w1 (attach_wiring w2 callee)" + unfolding attach_wiring_def comp_wiring_def + by (simp add: split_def map_fun_def o_def map_gpv_conv_map_gpv' map_gpv'_comp id_def map_prod_def) + + + + +subsection \Probabilistic finite converter\ + +coinductive pfinite_converter :: "('a, 'b) \ \ ('c, 'd) \ \ ('a, 'b, 'c, 'd) converter \ bool" + for \ \' where + pfinite_converterI: "pfinite_converter \ \' conv" if + "\a. a \ outs_\ \ \ pfinite_gpv \' (run_converter conv a)" + "\a b conv'. \ a \ outs_\ \; (b, conv') \ results_gpv \' (run_converter conv a) \ \ pfinite_converter \ \' conv'" + +lemma pfinite_converter_coinduct[consumes 1, case_names pfinite_converter, case_conclusion pfinite_converter pfinite step, coinduct pred: pfinite_converter]: + assumes "X conv" + and step: "\conv a. \ X conv; a \ outs_\ \ \ \ pfinite_gpv \' (run_converter conv a) \ + (\(b, conv') \ results_gpv \' (run_converter conv a). X conv' \ pfinite_converter \ \' conv')" + shows "pfinite_converter \ \' conv" + using assms(1) by(rule pfinite_converter.coinduct)(auto dest: step) + +lemma pfinite_converterD: + "\ pfinite_converter \ \' conv; a \ outs_\ \ \ + \ pfinite_gpv \' (run_converter conv a) \ + (\(b, conv') \ results_gpv \' (run_converter conv a). pfinite_converter \ \' conv')" + by(auto elim: pfinite_converter.cases) + +lemma pfinite_converter_bot1 [simp]: "pfinite_converter bot \ conv" + by(rule pfinite_converterI) auto + +lemma pfinite_converter_mono: + assumes *: "pfinite_converter \1 \2 conv" + and le: "outs_\ \1' \ outs_\ \1" "\2 \ \2'" + and WT: "\1, \2 \\<^sub>C conv \" + shows "pfinite_converter \1' \2' conv" + using * WT + apply(coinduction arbitrary: conv) + apply(drule pfinite_converterD) + apply(erule le(1)[THEN subsetD]) + apply(drule WT_converterD') + apply(erule le(1)[THEN subsetD]) + using le(2)[THEN responses_\_mono] + by(auto intro: pfinite_gpv_mono[OF _ le(2)] results_gpv_mono[OF le(2), THEN subsetD] dest: bspec) + +context raw_converter_invariant begin +lemma pfinite_converter_of_callee: + assumes step: "\x s. \ x \ outs_\ \; I s \ \ pfinite_gpv \' (callee s x)" + and I: "I s" + shows "pfinite_converter \ \' (converter_of_callee callee s)" + using I + by(coinduction arbitrary: s)(auto 4 3 simp add: step dest: results_callee) +end + +lemma raw_converter_invariant_run_pfinite_converter: + "raw_converter_invariant \ \' run_converter (\conv. pfinite_converter \ \' conv \ \,\' \\<^sub>C conv \)" + by(unfold_locales)(auto dest: WT_converterD pfinite_converterD) + +interpretation run_pfinite_converter: raw_converter_invariant + \ \' run_converter "\conv. pfinite_converter \ \' conv \ \,\' \\<^sub>C conv \" for \ \' + by(rule raw_converter_invariant_run_pfinite_converter) + +named_theorems pfinite_intro "Introduction rules for probabilistic finiteness" + +lemma pfinite_id_converter [pfinite_intro]: "pfinite_converter \ \ id_converter" + by(coinduction) simp + +lemma pfinite_fail_converter [pfinite_intro]: "pfinite_converter \ \' fail_converter" + by coinduction simp + +lemma pfinite_parallel_converter2 [pfinite_intro]: + "pfinite_converter (\1 \\<^sub>\ \2) (\1' \\<^sub>\ \2') (conv1 |\<^sub>= conv2)" + if "pfinite_converter \1 \1' conv1" "pfinite_converter \2 \2' conv2" + using that by(coinduction arbitrary: conv1 conv2)(fastforce dest: pfinite_converterD) + +context raw_converter_invariant begin + +lemma expectation_gpv_1_le_inline: + defines "expectation_gpv2 \ expectation_gpv 1 \'" + assumes callee: "\s x. \ x \ outs_\ \; I s \ \ pfinite_gpv \' (callee s x)" + and WT_gpv: "\ \g gpv \" + and I: "I s" + and f_le_1: "\x. f x \ 1" + shows "expectation_gpv 1 \ f gpv \ expectation_gpv2 (\(x, s). f x) (inline callee gpv s)" + using WT_gpv I +proof(induction arbitrary: gpv s rule: expectation_gpv_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step expectation_gpv') + have "(\\<^sup>+ x. (case x of Pure a \ f a | IO out c \ \r\responses_\ \ out. expectation_gpv' (c r)) \measure_spmf (the_gpv gpv)) + 1 * ennreal (pmf (the_gpv gpv) None) = + (\\<^sup>+ x. pmf (the_gpv gpv) x * (case x of Some (Pure a) \ f a | Some (IO out c) \ \r\responses_\ \ out. expectation_gpv' (c r) | None \ 1))" + apply(simp add: nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space nn_integral_measure_pmf) + apply(subst (2) nn_integral_disjoint_pair_countspace[where B="range Some" and C="{None}", simplified, folded UNIV_option_conv, simplified]) + apply(auto simp add: mult.commute intro!: nn_integral_cong split: split_indicator) + done + also have "\ \ (\\<^sup>+ x. pmf (the_gpv gpv) x * (case x of None \ 1 | Some (Pure a) \ f a | Some (IO out c) \ + (\\<^sup>+ x. ennreal (pmf (the_gpv (callee s out) \ case_generat (\(x, y). inline1 callee (c x) y) (\out rpv'. return_spmf (Inr (out, rpv', c)))) x) * + (case x of None \ 1 | Some (Inl (a, s)) \ f a + | Some (Inr (r, rpv, rpv')) \ \x\responses_\ \' r. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))))" + (is "nn_integral _ ?lhs \ nn_integral _ ?rhs") + proof(rule nn_integral_mono) + fix x :: "('a, 'call, ('a, 'call, 'ret) rpv) generat option" + consider (None) "x = None" | (Pure) a where "x = Some (Pure a)" + | (IO) out c where "x = Some (IO out c)" "IO out c \ set_spmf (the_gpv gpv)" + | (outside) out c where "x = Some (IO out c)" "IO out c \ set_spmf (the_gpv gpv)" + by (metis dest_IO.elims not_None_eq) + then show "?lhs x \ ?rhs x" + proof cases + case None then show ?thesis by simp + next + case Pure then show ?thesis by simp + next + case (IO out c) + with step.prems have out: "out \ outs_\ \" by(auto dest: WT_gpvD) + then obtain response where resp: "response \ responses_\ \ out" unfolding in_outs_\_iff_responses_\ by blast + with out step.prems IO have WT_resp [WT_intro]: "\ \g c response \" by(auto dest: WT_gpvD) + have exp_resp: "expectation_gpv' (c response) \ 1" + using step.hyps[of "c response"] expectation_gpv_mono[of 1 1 f "\_. 1" \ "c response"] expectation_gpv_const_le[OF WT_resp, of 1 1] + by(simp add: le_fun_def f_le_1) + + have "(\r\responses_\ \ out. expectation_gpv' (c r)) = + (\\<^sup>+ generat. (\r\responses_\ \ out. expectation_gpv' (c r)) \measure_spmf (the_gpv (callee s out))) + + (\r\responses_\ \ out. expectation_gpv' (c r)) * (1 - ennreal (weight_spmf (the_gpv (callee s out))))" + by(simp add: measure_spmf.emeasure_eq_measure add_mult_distrib2[symmetric] semiring_class.distrib_left[symmetric] add_diff_inverse_ennreal weight_spmf_le_1) + also have "\ \ (\\<^sup>+ generat. (\r\responses_\ \ out. expectation_gpv' (c r)) \measure_spmf (the_gpv (callee s out))) + + 1 * ennreal (pmf (the_gpv (callee s out)) None)" unfolding pmf_None_eq_weight_spmf + by(intro add_mono mult_mono order_refl INF_lower2[OF resp])(auto simp add: ennreal_minus[symmetric] weight_spmf_le_1 exp_resp) + also have "\ = (\\<^sup>+ z. ennreal (pmf (the_gpv (callee s out)) z) * (case z of None \ 1 | Some generat \ (\r\responses_\ \ out. expectation_gpv' (c r))))" + apply(simp add: nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space nn_integral_measure_pmf del: nn_integral_const) + apply(subst (2) nn_integral_disjoint_pair_countspace[where B="range Some" and C="{None}", simplified, folded UNIV_option_conv, simplified]) + apply(auto simp add: mult.commute intro!: nn_integral_cong split: split_indicator) + done + also have "\ \ (\\<^sup>+ z. ennreal (pmf (the_gpv (callee s out)) z) * + (case z of None \ 1 | Some (IO out' rpv') \ \x\responses_\ \' out'. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (c x) s')) (rpv' x) + | Some (Pure (r, s')) \ (\\<^sup>+ x. ennreal (pmf (inline1 callee (c r) s') x) * (case x of None \ 1 | Some (Inl (a, s)) \ f a | Some (Inr (out', rpv, rpv')) \ + \x\responses_\ \' out'. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))))" + (is "nn_integral _ ?lhs2 \ nn_integral _ ?rhs2") + proof(intro nn_integral_mono) + fix z :: "('ret \ 's, 'call', ('ret \ 's, 'call', 'ret') rpv) generat option" + consider (None) "z = None" | (Pure) x' s' where "z = Some (Pure (x', s'))" "Pure (x', s') \ set_spmf (the_gpv (callee s out))" + | (IO') out' c' where "z = Some (IO out' c')" "IO out' c' \ set_spmf (the_gpv (callee s out))" + | (outside) generat where "z = Some generat" "generat \ set_spmf (the_gpv (callee s out))" + by (metis dest_IO.elims not_Some_eq old.prod.exhaust) + then show "?lhs2 z \ ?rhs2 z" + proof cases + case None then show ?thesis by simp + next + case Pure + hence "(x', s') \ results_gpv \' (callee s out)" by(simp add: results_gpv.Pure) + with results_callee step.prems out have x: "x' \ responses_\ \ out" and s': "I s'" by auto + with IO out step.prems have WT_c [WT_intro]: "\ \g c x' \" by(auto dest: WT_gpvD) + from x have "(INF r\responses_\ \ out. expectation_gpv' (c r)) \ expectation_gpv' (c x')" by(rule INF_lower) + also have "\ \ expectation_gpv2 (\(x, s). f x) (inline callee (c x') s')" using WT_c s' by(rule step.IH) + also have "\ = \\<^sup>+ xx. (case xx of Inl (x, _) \ f x + | Inr (out', callee', rpv) \ INF r'\responses_\ \' out'. expectation_gpv 1 \' (\(r, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv r) s')) (callee' r')) + \measure_spmf (inline1 callee (c x') s') + ennreal (pmf (the_gpv (inline callee (c x') s')) None)" + unfolding expectation_gpv2_def + by(subst expectation_gpv.simps)(auto simp add: inline_sel split_def o_def intro!: nn_integral_cong split: generat.split sum.split) + also have "\ = (\\<^sup>+ xx. ennreal (pmf (inline1 callee (c x') s') xx) * (case xx of None \ 1 | Some (Inl (x, _)) \ f x + | Some (Inr (out', callee', rpv)) \ INF r'\responses_\ \' out'. expectation_gpv 1 \' (\(r, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv r) s')) (callee' r')))" + apply(subst inline_sel) + apply(simp add: nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space nn_integral_measure_pmf pmf_map_spmf_None del: nn_integral_const) + apply(subst (2) nn_integral_disjoint_pair_countspace[where B="range Some" and C="{None}", simplified, folded UNIV_option_conv, simplified]) + apply(auto simp add: mult.commute intro!: nn_integral_cong split: split_indicator) + done + finally show ?thesis using Pure by(simp add: mult_mono) + next + case IO' + then have out': "out' \ outs_\ \'" using WT_callee out step.prems by(auto dest: WT_gpvD) + have "(INF r\responses_\ \ out. expectation_gpv' (c r)) \ min (INF (r, s')\(\r'\responses_\ \' out'. results_gpv \' (c' r')). expectation_gpv' (c r)) 1" + using IO' results_callee[OF out, of s] step.prems by(intro INF_mono min.boundedI)(auto intro: results_gpv.IO intro!: INF_lower2[OF resp] exp_resp) + also have "\ \ (INF r'\responses_\ \' out'. min (INF (r, s')\results_gpv \' (c' r'). expectation_gpv' (c r)) 1)" + using resp out' unfolding inf_min[symmetric] in_outs_\_iff_responses_\ + by(subst INF_inf_const2)(auto simp add: INF_UNION) + also have "\ \ (INF r'\responses_\ \' out'. expectation_gpv 1 \' (\(r', s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (c r') s')) (c' r'))" + (is "_ \ (INF r'\_. ?r r')") + proof(rule INF_mono, rule bexI) + fix r' + assume r': "r' \ responses_\ \' out'" + have fin: "pfinite_gpv \' (c' r')" using callee[OF out, of s] IO' r' WT_callee[OF out, of s] step.prems by(auto dest: pfinite_gpv_ContD) + have "min (INF (r, s')\results_gpv \' (c' r'). expectation_gpv' (c r)) 1 \ min (INF (r, s')\results_gpv \' (c' r'). expectation_gpv2 (\(x, s). f x) (inline callee (c r) s')) 1" + using IO IO' step.prems out results_callee[OF out, of s] r' + by(intro min.mono)(auto intro!: INF_mono rev_bexI step.IH dest: WT_gpv_ContD intro: results_gpv.IO) + also have "\ \ ?r r'" unfolding expectation_gpv2_def using fin by(rule pfinite_INF_le_expectation_gpv) + finally show "min (INF (r, s')\results_gpv \' (c' r'). expectation_gpv' (c r)) 1 \ \" . + qed + finally show ?thesis using IO' by(simp add: mult_mono) + next + case outside then show ?thesis by(simp add: in_set_spmf_iff_spmf) + qed + qed + also have "\ = (\\<^sup>+ z. \\<^sup>+ x. + ennreal (pmf (the_gpv (callee s out)) z) * + ennreal (pmf (case z of None \ return_pmf None | Some (Pure (x, xb)) \ inline1 callee (c x) xb | Some (IO out rpv') \ return_spmf (Inr (out, rpv', c))) x) * + (case x of None \ 1 | Some (Inl (a, s)) \ f a | Some (Inr (out, rpv, rpv')) \ \x\responses_\ \' out. + expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))" + (is "\ = (\\<^sup>+ z. \\<^sup>+ x. ?f x z)") + by(auto intro!: nn_integral_cong split!: option.split generat.split simp add: mult.assoc nn_integral_cmult ennreal_indicator) + also have "(\\<^sup>+ z. \\<^sup>+ x. ?f x z) = (\\<^sup>+ x. \\<^sup>+ z. ?f x z)" + by(subst nn_integral_fst_count_space[where f="case_prod _", simplified])(simp add: nn_integral_snd_count_space[symmetric]) + also have "\ = (\\<^sup>+ x. + ennreal (pmf (the_gpv (callee s out) \ case_generat (\(x, y). inline1 callee (c x) y) (\out rpv'. return_spmf (Inr (out, rpv', c)))) x) * + (case x of None \ 1 | Some (Inl (a, s)) \ f a | Some (Inr (r, rpv, rpv')) \ + \x\responses_\ \' r. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))" + by(simp add: bind_spmf_def ennreal_pmf_bind nn_integral_multc[symmetric] nn_integral_measure_pmf) + finally show ?thesis using IO by(auto intro!: mult_mono) + next + case outside then show ?thesis by(simp add: in_set_spmf_iff_spmf) + qed + qed + also have "\ = (\\<^sup>+ y. \\<^sup>+ x. + ennreal (pmf (the_gpv gpv) y) * + ennreal (case y of None \ pmf (return_pmf None) x | Some (Pure xa) \ pmf (return_spmf (Inl (xa, s))) x + | Some (IO out rpv) \ + pmf (bind_spmf (the_gpv (callee s out)) (\generat' \ case generat' of Pure (x, y) \ inline1 callee (rpv x) y | IO out rpv' \ return_spmf (Inr (out, rpv', rpv)))) x) * + (case x of None \ 1 | Some (Inl (a, s)) \ f a + | Some (Inr (out, rpv, rpv')) \ \x\responses_\ \' out. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))" + (is "_ = (\\<^sup>+ y. \\<^sup>+ x. ?f x y)") + by(auto intro!: nn_integral_cong split!: option.split generat.split simp add: nn_integral_cmult mult.assoc ennreal_indicator) + also have "(\\<^sup>+ y. \\<^sup>+ x. ?f x y) = (\\<^sup>+ x. \\<^sup>+ y. ?f x y)" + by(subst nn_integral_fst_count_space[where f="case_prod _", simplified])(simp add: nn_integral_snd_count_space[symmetric]) + also have "\ = (\\<^sup>+ x. (pmf (inline1 callee gpv s) x) * (case x of None \ 1 | Some (Inl (a, s)) \ f a | + Some (Inr (out, rpv, rpv')) \ \x\responses_\ \' out. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)))" + by(rewrite in "_ = \" inline1.simps) + (auto simp add: bind_spmf_def ennreal_pmf_bind nn_integral_multc[symmetric] nn_integral_measure_pmf intro!: nn_integral_cong split: option.split generat.split) + also have "\ = (\\<^sup>+ res. (case res of Inl (a, s) \ f a + | Inr (out, rpv, rpv') \ \x\responses_\ \' out. expectation_gpv 1 \' (\(x, s'). expectation_gpv 1 \' (\(x, s). f x) (inline callee (rpv' x) s')) (rpv x)) + \measure_spmf (inline1 callee gpv s) + + ennreal (pmf (inline1 callee gpv s) None))" + apply(simp add: nn_integral_measure_spmf_conv_measure_pmf nn_integral_restrict_space nn_integral_measure_pmf) + apply(subst nn_integral_disjoint_pair_countspace[where B="range Some" and C="{None}", simplified, folded UNIV_option_conv, simplified]) + apply(auto simp add: mult.commute intro!: nn_integral_cong split: split_indicator) + done + also have "\ = expectation_gpv2 (\(x, s). f x) (inline callee gpv s)" unfolding expectation_gpv2_def + by(rewrite in "_ = \" expectation_gpv.simps, subst (1 2) inline_sel) + (simp add: o_def pmf_map_spmf_None sum.case_distrib[where h="case_generat _ _"] split_def cong: sum.case_cong) + finally show ?case . +qed + +lemma pfinite_inline: + assumes fin: "pfinite_gpv \ gpv" + and WT: "\ \g gpv \" + and callee: "\s x. \ x \ outs_\ \; I s \ \ pfinite_gpv \' (callee s x)" + and I: "I s" + shows "pfinite_gpv \' (inline callee gpv s)" +unfolding pgen_lossless_gpv_def +proof(rule antisym) + have WT': "\' \g inline callee gpv s \" using WT I by(rule WT_gpv_inline_invar) + from expectation_gpv_const_le[OF WT', of 1 1] + show "expectation_gpv 1 \' (\_. 1) (inline callee gpv s) \ 1" by(simp add: max_def) + + have "1 = expectation_gpv 1 \ (\_. 1) gpv" using fin by(simp add: pgen_lossless_gpv_def) + also have "\ \ expectation_gpv 1 \' (\_. 1) (inline callee gpv s)" + by(rule expectation_gpv_1_le_inline[unfolded split_def]; rule callee I WT WT_callee order_refl) + finally show "1 \ \" . +qed + +end + +lemma pfinite_comp_converter [pfinite_intro]: + "pfinite_converter \1 \3 (conv1 \ conv2)" + if "pfinite_converter \1 \2 conv1" "pfinite_converter \2 \3 conv2" "\1,\2 \\<^sub>C conv1 \" "\2,\3 \\<^sub>C conv2 \" + using that +proof(coinduction arbitrary: conv1 conv2) + case pfinite_converter + have conv1: "pfinite_gpv \2 (run_converter conv1 a)" + using pfinite_converter(1, 5) by(simp add: pfinite_converterD) + have conv2: "\2 \g run_converter conv1 a \" + using pfinite_converter(3, 5) by(simp add: WT_converterD) + have ?pfinite using pfinite_converter(2,4,5) + by(auto intro!:run_pfinite_converter.pfinite_inline[OF conv1] dest: pfinite_converterD intro: conv2) + moreover have ?step (is "\(b, conv')\?res. ?P b conv' \ _") + proof(clarify) + fix b conv'' + assume "(b, conv'') \ ?res" + then obtain conv1' conv2' where [simp]: "conv'' = comp_converter conv1' conv2'" + and inline: "((b, conv1'), conv2') \ results_gpv \3 (inline run_converter (run_converter conv1 a) conv2)" + by auto + from run_pfinite_converter.results_gpv_inline[OF inline conv2] pfinite_converter(2,4) + have run: "(b, conv1') \ results_gpv \2 (run_converter conv1 a)" + and *: "pfinite_converter \2 \3 conv2'" "\2, \3 \\<^sub>C conv2' \" by auto + with WT_converterD(2)[OF pfinite_converter(3,5) run] pfinite_converterD[THEN conjunct2, rule_format, OF pfinite_converter(1,5) run] + show "?P b conv''" by auto + qed + ultimately show ?case .. +qed + +lemma pfinite_map_converter [pfinite_intro]: + "pfinite_converter \ \' (map_converter f g f' g' conv)" if + *: "pfinite_converter (map_\ (inv_into UNIV f) (inv_into UNIV g) \) (map_\ f' g' \') conv" + and f: "inj f" and g: "surj g" + using * +proof(coinduction arbitrary: conv) + case (pfinite_converter a conv) + with f have a: "inv_into UNIV f (f a) \ outs_\ \" by simp + with pfinite_converterD[OF \pfinite_converter _ _ conv\, of "f a"] have "?pfinite" by simp + moreover have ?step + proof(safe) + fix r conv' + assume "(r, conv') \ results_gpv \' (run_converter (map_converter f g f' g' conv) a)" + then obtain r' conv'' + where results: "(r', conv'') \ results_gpv (map_\ f' g' \') (run_converter conv (f a))" + and r: "r = g r'" + and conv': "conv' = map_converter f g f' g' conv''" + by auto + from pfinite_converterD[OF \pfinite_converter _ _ conv\, THEN conjunct2, rule_format, OF _ results] a r conv' + show "\conv. conv' = map_converter f g f' g' conv \ + pfinite_converter (map_\ (inv_into UNIV f) (inv_into UNIV g) \) (map_\ f' g' \') conv" + by auto + qed + ultimately show ?case .. +qed + +lemma pfinite_lassocr\<^sub>C [pfinite_intro]: "pfinite_converter ((\1 \\<^sub>\ \2) \\<^sub>\ \3) (\1 \\<^sub>\ (\2 \\<^sub>\ \3)) lassocr\<^sub>C" + by(coinduction)(auto simp add: lassocr\<^sub>C_def) + +lemma pfinite_rassocl\<^sub>C [pfinite_intro]: "pfinite_converter (\1 \\<^sub>\ (\2 \\<^sub>\ \3)) ((\1 \\<^sub>\ \2) \\<^sub>\ \3) rassocl\<^sub>C" + by(coinduction)(auto simp add: rassocl\<^sub>C_def) + +lemma pfinite_swap\<^sub>C [pfinite_intro]: "pfinite_converter (\1 \\<^sub>\ \2) (\2 \\<^sub>\ \1) swap\<^sub>C" + by(coinduction)(auto simp add: swap\<^sub>C_def) + +lemma pfinite_swap_lassocr [pfinite_intro]: "pfinite_converter (\1 \\<^sub>\ (\2 \\<^sub>\ \3)) (\2 \\<^sub>\ (\1 \\<^sub>\ \3)) swap_lassocr" + unfolding swap_lassocr_def by(rule pfinite_intro WT_intro)+ + +lemma pfinite_swap_rassocl [pfinite_intro]: "pfinite_converter ((\1 \\<^sub>\ \2) \\<^sub>\ \3) ((\1 \\<^sub>\ \3) \\<^sub>\ \2) swap_rassocl" + unfolding swap_rassocl_def by(rule pfinite_intro WT_intro)+ + +lemma pfinite_parallel_wiring [pfinite_intro]: + "pfinite_converter ((\1 \\<^sub>\ \2) \\<^sub>\ (\3 \\<^sub>\ \4)) ((\1 \\<^sub>\ \3) \\<^sub>\ (\2 \\<^sub>\ \4)) parallel_wiring" + unfolding parallel_wiring_def by(rule pfinite_intro WT_intro)+ + +lemma pfinite_parallel_converter [pfinite_intro]: + "pfinite_converter (\1 \\<^sub>\ \2) \3 (conv1 |\<^sub>\ conv2)" + if "pfinite_converter \1 \3 conv1" and "pfinite_converter \2 \3 conv2" + using that by(coinduction arbitrary: conv1 conv2)(fastforce dest: pfinite_converterD) + +lemma pfinite_converter_of_resource [simp, pfinite_intro]: "pfinite_converter \1 \2 (converter_of_resource res)" + by(coinduction arbitrary: res) auto + +subsection \colossless converter\ + +coinductive colossless_converter :: "('a, 'b) \ \ ('c, 'd) \ \ ('a, 'b, 'c, 'd) converter \ bool" + for \ \' where + colossless_converterI: + "colossless_converter \ \' conv" if + "\a. a \ outs_\ \ \ colossless_gpv \' (run_converter conv a)" + "\a b conv'. \ a \ outs_\ \; (b, conv') \ results_gpv \' (run_converter conv a) \ \ colossless_converter \ \' conv'" + +lemma colossless_converter_coinduct[consumes 1, case_names colossless_converter, case_conclusion colossless_converter plossless step, coinduct pred: colossless_converter]: + assumes "X conv" + and step: "\conv a. \ X conv; a \ outs_\ \ \ \ colossless_gpv \' (run_converter conv a) \ + (\(b, conv') \ results_gpv \' (run_converter conv a). X conv' \ colossless_converter \ \' conv')" + shows "colossless_converter \ \' conv" + using assms(1) by(rule colossless_converter.coinduct)(auto dest: step) + +lemma colossless_converterD: + "\ colossless_converter \ \' conv; a \ outs_\ \ \ + \ colossless_gpv \' (run_converter conv a) \ + (\(b, conv') \ results_gpv \' (run_converter conv a). colossless_converter \ \' conv')" + by(auto elim: colossless_converter.cases) + +lemma colossless_converter_bot1 [simp]: "colossless_converter bot \ conv" + by(rule colossless_converterI) auto + +lemma raw_converter_invariant_run_colossless_converter: "raw_converter_invariant \ \' run_converter (\conv. colossless_converter \ \' conv \ \,\' \\<^sub>C conv \)" + by(unfold_locales)(auto dest: WT_converterD colossless_converterD) + +interpretation run_colossless_converter: raw_converter_invariant + \ \' run_converter "\conv. colossless_converter \ \' conv \ \,\' \\<^sub>C conv \" for \ \' + by(rule raw_converter_invariant_run_colossless_converter) + +lemma colossless_const_converter [simp]: "colossless_converter \ \' (const_converter x)" + by(coinduction)(auto) + +subsection \trace equivalence\ + +lemma distinguish_trace_eq: (* generalized from Distinguisher.thy *) + assumes distinguish: "\distinguisher. \ \g distinguisher \ \ connect distinguisher res = connect distinguisher res'" + shows "outs_\ \ \\<^sub>R res \ res'" + using assms by(rule distinguish_trace_eq)(auto intro: WT_fail_resource) + +lemma attach_trace_eq': + assumes eq: "outs_\ \ \\<^sub>R res1 \ res2" + and WT1 [WT_intro]: "\ \res res1 \" + and WT2 [WT_intro]: "\ \res res2 \" + and WT_conv [WT_intro]: "\',\ \\<^sub>C conv \" + shows "outs_\ \' \\<^sub>R conv \ res1 \ conv \ res2" +proof(rule distinguish_trace_eq) + fix \ :: "('c, 'd) distinguisher" + assume [WT_intro]: "\' \g \ \" + have "connect (absorb \ conv) res1 = connect (absorb \ conv) res2" using eq + by(rule connect_cong_trace)(rule WT_intro | fold WT_gpv_iff_outs_gpv)+ + then show "connect \ (conv \ res1) = connect \ (conv \ res2)" by(simp add: distinguish_attach) +qed + +lemma trace_callee_eq_trans [trans]: + "\ trace_callee_eq callee1 callee2 A p q; trace_callee_eq callee2 callee3 A q r \ + \ trace_callee_eq callee1 callee3 A p r" + by(simp add: trace_callee_eq_def) + +lemma trace_eq'_parallel_resource: + fixes res1 :: "('a, 'b) resource" and res2 :: "('c, 'd) resource" + assumes 1: "trace_eq' A res1 res1'" + and 2: "trace_eq' B res2 res2'" + shows "trace_eq' (A <+> B) (res1 \ res2) (res1' \ res2')" +proof - + let ?\ = "\_uniform A (UNIV :: 'b set) \\<^sub>\ \_uniform B (UNIV :: 'd set)" + have "trace_eq' (outs_\ ?\) (res1 \ res2) (res1' \ res2)" + apply(subst (1 2) attach_converter_of_resource_conv_parallel_resource2[symmetric]) + apply(rule attach_trace_eq'[where ?\ = "\_uniform A UNIV"]; auto simp add: 1 intro: WT_intro WT_resource_\_uniform_UNIV) + done + also have "trace_eq' (outs_\ ?\) (res1' \ res2) (res1' \ res2')" + apply(subst (1 2) attach_converter_of_resource_conv_parallel_resource[symmetric]) + apply(rule attach_trace_eq'[where ?\ = "\_uniform B UNIV"]; auto simp add: 2 intro: WT_intro WT_resource_\_uniform_UNIV) + done + finally show ?thesis by simp +qed + +proposition trace_callee_eq_coinduct [consumes 1, case_names step sim]: (* stronger version of trace'_eq_simI *) + fixes callee1 :: "('a, 'b, 's1) callee" and callee2 :: "('a, 'b, 's2) callee" + assumes start: "S p q" + and step: "\p q a. \ S p q; a \ A \ \ + bind_spmf p (\s. map_spmf fst (callee1 s a)) = bind_spmf q (\s. map_spmf fst (callee2 s a))" + and sim: "\p q a res res' b s'' s'. \ S p q; a \ A; res \ set_spmf p; (b, s'') \ set_spmf (callee1 res a); res' \ set_spmf q; (b, s') \ set_spmf (callee2 res' a) \ + \ S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" + shows "trace_callee_eq callee1 callee2 A p q" +proof(rule trace_callee_eqI) + fix xs :: "('a \ 'b) list" and z + assume xs: "set xs \ A \ UNIV" and z: "z \ A" + + from start show "trace_callee callee1 p xs z = trace_callee callee2 q xs z" using xs + proof(induction xs arbitrary: p q) + case Nil + then show ?case using z by(simp add: step) + next + case (Cons xy xs) + obtain x y where xy [simp]: "xy = (x, y)" by(cases xy) + have "trace_callee callee1 p (xy # xs) z = + trace_callee callee1 (cond_spmf_fst (bind_spmf p (\s. callee1 s x)) y) xs z" + by(simp add: bind_map_spmf split_def o_def) + also have "\ = trace_callee callee2 (cond_spmf_fst (bind_spmf q (\s. callee2 s x)) y) xs z" + proof(cases "\s \ set_spmf q. \s'. (y, s') \ set_spmf (callee2 s x)") + case True + then obtain s s' where ss': "s \ set_spmf q" "(y, s') \ set_spmf (callee2 s x)" by blast + from Cons have "x \ A" by simp + from ss' step[THEN arg_cong[where f="set_spmf"], OF \S p q\ this] obtain ss ss' + where "ss \ set_spmf p" "(y, ss') \ set_spmf (callee1 ss x)" + by(clarsimp simp add: bind_UNION) force + from sim[OF \S p q\ _ this ss'] show ?thesis using Cons.prems by (auto intro: Cons.IH) + next + case False + then have "cond_spmf_fst (bind_spmf q (\s. callee2 s x)) y = return_pmf None" + by(auto simp add: bind_eq_return_pmf_None) + moreover from step[OF \S p q\, of x, THEN arg_cong[where f=set_spmf], THEN eq_refl] Cons.prems False + have "cond_spmf_fst (bind_spmf p (\s. callee1 s x)) y = return_pmf None" + by(clarsimp simp add: bind_eq_return_pmf_None)(rule ccontr; fastforce) + ultimately show ?thesis by(simp del: cond_spmf_fst_eq_return_None) + qed + also have "\ = trace_callee callee2 q (xy # xs) z" + by(simp add: split_def o_def) + finally show ?case . + qed +qed + +proposition trace_callee_eq_coinduct_strong [consumes 1, case_names step sim, case_conclusion step lhs rhs, case_conclusion sim sim eq]: + fixes callee1 :: "('a, 'b, 's1) callee" and callee2 :: "('a, 'b, 's2) callee" + assumes start: "S p q" + and step: "\p q a. \ S p q; a \ A \ \ + bind_spmf p (\s. map_spmf fst (callee1 s a)) = bind_spmf q (\s. map_spmf fst (callee2 s a))" + and sim: "\p q a res res' b s'' s'. \ S p q; a \ A; res \ set_spmf p; (b, s'') \ set_spmf (callee1 res a); res' \ set_spmf q; (b, s') \ set_spmf (callee2 res' a) \ + \ S (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) + (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b) \ + trace_callee_eq callee1 callee2 A (cond_spmf_fst (bind_spmf p (\s. callee1 s a)) b) (cond_spmf_fst (bind_spmf q (\s. callee2 s a)) b)" + shows "trace_callee_eq callee1 callee2 A p q" +proof - + from start have "S p q \ trace_callee_eq callee1 callee2 A p q" by simp + thus ?thesis + apply(rule trace_callee_eq_coinduct) + apply(erule disjE) + apply(erule (1) step) + apply(drule trace_callee_eqD[where xs="[]"]; simp) + apply(erule disjE) + apply(erule (5) sim) + apply(rule disjI2) + apply(rule trace_callee_eqI) + apply(drule trace_callee_eqD[where xs="(_, _) # _"]) + apply simp_all + done +qed + +lemma trace_callee_return_pmf_None [simp]: + "trace_callee_eq callee1 callee2 A (return_pmf None) (return_pmf None)" + by(rule trace_callee_eqI) simp + +lemma trace_callee_eq_sym [sym]: "trace_callee_eq callee1 callee2 A p q \ trace_callee_eq callee2 callee1 A q p" + by(simp add: trace_callee_eq_def) + +lemma eq_resource_on_imp_trace_eq: "A \\<^sub>R res1 \ res2" if "A \\<^sub>R res1 \ res2" +proof - + have "outs_\ (\_uniform A UNIV :: ('a, 'b) \) \\<^sub>R res1 \ res2" using that + by -(rule distinguish_trace_eq[OF connect_eq_resource_cong], simp+) + thus ?thesis by simp +qed + +lemma advantage_nonneg: "0 \ advantage \ res1 res2" + by(simp add: advantage_def) + +lemma comp_converter_of_resource_conv_parallel_converter: + "(converter_of_resource res |\<^sub>\ 1\<^sub>C) \ conv = converter_of_resource res |\<^sub>\ conv" + by(coinduction arbitrary: res conv) + (auto 4 3 simp add: rel_fun_def gpv.map_comp map_lift_spmf spmf_rel_map split_def map_gpv_conv_bind[symmetric] id_def[symmetric] gpv.rel_map split!: sum.split intro!: rel_spmf_reflI gpv.rel_refl_strong) + +lemma comp_converter_of_resource_conv_parallel_converter2: + "(1\<^sub>C |\<^sub>\ converter_of_resource res) \ conv = conv |\<^sub>\ converter_of_resource res" + by(coinduction arbitrary: res conv) + (auto 4 3 simp add: rel_fun_def gpv.map_comp map_lift_spmf spmf_rel_map split_def map_gpv_conv_bind[symmetric] id_def[symmetric] gpv.rel_map split!: sum.split intro!: rel_spmf_reflI gpv.rel_refl_strong) + +lemma parallel_converter_map_converter: + "map_converter f g f' g' conv1 |\<^sub>\ map_converter f'' g'' f' g' conv2 = + map_converter (map_sum f f'') (map_sum g g'') f' g' (conv1 |\<^sub>\ conv2)" + using parallel_callee_parametric[ + where A="conversep (BNF_Def.Grp UNIV f)" and B="BNF_Def.Grp UNIV g" and C="BNF_Def.Grp UNIV f'" and R="conversep (BNF_Def.Grp UNIV g')" and A'="conversep (BNF_Def.Grp UNIV f'')" and B'="BNF_Def.Grp UNIV g''", + unfolded rel_converter_Grp sum.rel_conversep sum.rel_Grp, + simplified, + unfolded rel_converter_Grp] + by(simp add: rel_fun_def Grp_def) + +lemma map_converter_parallel_converter_out2: + "conv1 |\<^sub>\ map_converter f g id id conv2 = map_converter (map_sum id f) (map_sum id g) id id (conv1 |\<^sub>\ conv2)" + by(rule parallel_converter_map_converter[where f=id and g=id and f'=id and g'=id, simplified]) + +lemma parallel_converter_assoc2: + "parallel_converter conv1 (parallel_converter conv2 conv3) = + map_converter lsumr rsuml id id (parallel_converter (parallel_converter conv1 conv2) conv3)" + by(coinduction arbitrary: conv1 conv2 conv3) + (auto 4 5 intro!: rel_funI gpv.rel_refl_strong split: sum.split simp add: gpv.rel_map map_gpv'_id map_gpv_conv_map_gpv'[symmetric]) + +lemma parallel_converter_of_resource: + "converter_of_resource res1 |\<^sub>\ converter_of_resource res2 = converter_of_resource (res1 \ res2)" + by(coinduction arbitrary: res1 res2) + (auto 4 3 simp add: rel_fun_def map_lift_spmf spmf_rel_map intro!: rel_spmf_reflI split!: sum.split) + +lemma map_Inr_parallel_converter: + "map_converter Inr f g h (conv1 |\<^sub>\ conv2) = map_converter id (f \ Inr) g h conv2" + (is "?lhs = ?rhs") +proof - + have "?lhs = map_converter Inr f id id (map_converter id id g h conv1 |\<^sub>\ map_converter id id g h conv2)" + by(simp add: parallel_converter_map_converter sum.map_id0) + also have "map_converter Inr f id id (conv1 |\<^sub>\ conv2) = map_converter id (f \ Inr) id id conv2" for conv1 conv2 + by(coinduction arbitrary: conv2) + (auto simp add: rel_fun_def map_gpv_conv_map_gpv'[symmetric] gpv.rel_map intro!: gpv.rel_refl_strong) + also have "map_converter id (f \ Inr) id id (map_converter id id g h conv2) = ?rhs" by simp + finally show ?thesis . +qed + +lemma map_Inl_parallel_converter: + "map_converter Inl f g h (conv1 |\<^sub>\ conv2) = map_converter id (f \ Inl) g h conv1" + (is "?lhs = ?rhs") +proof - + have "?lhs = map_converter Inl f id id (map_converter id id g h conv1 |\<^sub>\ map_converter id id g h conv2)" + by(simp add: parallel_converter_map_converter sum.map_id0) + also have "map_converter Inl f id id (conv1 |\<^sub>\ conv2) = map_converter id (f \ Inl) id id conv1" for conv1 conv2 + by(coinduction arbitrary: conv1) + (auto simp add: rel_fun_def map_gpv_conv_map_gpv'[symmetric] gpv.rel_map intro!: gpv.rel_refl_strong) + also have "map_converter id (f \ Inl) id id (map_converter id id g h conv1) = ?rhs" by simp + finally show ?thesis . +qed + +lemma left_interface_parallel_converter: + "left_interface (conv1 |\<^sub>\ conv2) = left_interface conv1 |\<^sub>\ left_interface conv2" + by(coinduction arbitrary: conv1 conv2) + (auto 4 3 split!: sum.split simp add: rel_fun_def gpv.rel_map left_gpv_map[where h=id] sum.map_id0 intro!: gpv.rel_refl_strong) + +lemma right_interface_parallel_converter: + "right_interface (conv1 |\<^sub>\ conv2) = right_interface conv1 |\<^sub>\ right_interface conv2" + by(coinduction arbitrary: conv1 conv2) + (auto 4 3 split!: sum.split simp add: rel_fun_def gpv.rel_map right_gpv_map[where h=id] sum.map_id0 intro!: gpv.rel_refl_strong) + +lemma left_interface_converter_of_resource [simp]: + "left_interface (converter_of_resource res) = converter_of_resource res" + by(coinduction arbitrary: res)(auto simp add: rel_fun_def map_lift_spmf spmf_rel_map intro!: rel_spmf_reflI) + +lemma right_interface_converter_of_resource [simp]: + "right_interface (converter_of_resource res) = converter_of_resource res" + by(coinduction arbitrary: res)(auto simp add: rel_fun_def map_lift_spmf spmf_rel_map intro!: rel_spmf_reflI) + +lemma parallel_converter_swap: "map_converter swap_sum swap_sum id id (conv1 |\<^sub>\ conv2) = conv2 |\<^sub>\ conv1" + by(coinduction arbitrary: conv1 conv2) + (auto 4 3 split!: sum.split simp add: rel_fun_def map_gpv_conv_map_gpv'[symmetric] gpv.rel_map intro!: gpv.rel_refl_strong) + +lemma eq_\_converter_map_converter': + assumes "\'', map_\ f' g' \' \\<^sub>C conv1 \ conv2" + and "f ` outs_\ \ \ outs_\ \''" + and "\q\outs_\ \. g ` responses_\ \'' (f q) \ responses_\ \ q" + shows "\, \' \\<^sub>C map_converter f g f' g' conv1 \ map_converter f g f' g' conv2" + using assms(1) +proof(coinduction arbitrary: conv1 conv2) + case eq_\_converter + from this(2) have "f q \ outs_\ \''" using assms(2) by auto + from eq_\_converter(1)[THEN eq_\_converterD, OF this] eq_\_converter(2) + show ?case + apply simp + apply(rule eq_\_gpv_map_gpv') + apply(simp add: BNF_Def.vimage2p_def prod.rel_map) + apply(erule eq_\_gpv_mono') + using assms(3) + apply(auto 4 4 simp add: eq_onp_def) + done +qed + +lemma parallel_converter_eq_\_cong: + "\ \1, \ \\<^sub>C conv1 \ conv1'; \2, \ \\<^sub>C conv2 \ conv2' \ + \ \1 \\<^sub>\ \2, \ \\<^sub>C parallel_converter conv1 conv2 \ parallel_converter conv1' conv2'" + by(coinduction arbitrary: conv1 conv2 conv1' conv2') + (fastforce dest: eq_\_converterD elim!: eq_\_gpv_mono' simp add: eq_onp_def) + +\ \Helper lemmas for simplyfing @{term exec_gpv}\ +lemma + exec_gpv_parallel_oracle_right: + "exec_gpv (oracle1 \\<^sub>O oracle2) (right_gpv gpv) s = exec_gpv (\oracle2) gpv s" + unfolding spmf_rel_eq[symmetric] + apply (rule rel_spmf_mono) + by (rule exec_gpv_parametric'[where S="(=)" and A="(=)" and CALL="\l r. l = Inr r" and R="\l r. l = Inr r" , THEN rel_funD, THEN rel_funD, THEN rel_funD ]) + (auto simp add: prod.rel_eq extend_state_oracle_def parallel_oracle_def split_def + spmf_rel_map1 spmf_rel_map2 map_prod_def rel_spmf_reflI right_gpv_Inr_transfer intro!:rel_funI) + +lemma + exec_gpv_parallel_oracle_left: + "exec_gpv (oracle1 \\<^sub>O oracle2) (left_gpv gpv) s = exec_gpv (oracle1\) gpv s" (is "?L = ?R") + unfolding spmf_rel_eq[symmetric] + apply (rule rel_spmf_mono) + by (rule exec_gpv_parametric'[where S="(=)" and A="(=)" and CALL="\l r. l = Inl r" and R="\l r. l = Inl r" , THEN rel_funD, THEN rel_funD, THEN rel_funD ]) + (auto simp add: prod.rel_eq extend_state_oracle2_def parallel_oracle_def split_def + spmf_rel_map1 spmf_rel_map2 map_prod_def rel_spmf_reflI left_gpv_Inl_transfer intro!:rel_funI) + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/Observe_Failure.thy b/thys/Constructive_Cryptography_CM/Observe_Failure.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Observe_Failure.thy @@ -0,0 +1,579 @@ +theory Observe_Failure imports + More_CC +begin + +declare [[show_variants]] + +context fixes "oracle" :: "('s, 'in, 'out) oracle'" begin + +fun obsf_oracle :: "('s exception, 'in, 'out exception) oracle'" where + "obsf_oracle Fault x = return_spmf (Fault, Fault)" +| "obsf_oracle (OK s) x = TRY map_spmf (map_prod OK OK) (oracle s x) ELSE return_spmf (Fault, Fault)" + +end + +type_synonym ('a, 'b) resource_obsf = "('a, 'b exception) resource" + +translations + (type) "('a, 'b) resource_obsf" <= (type) "('a, 'b exception) resource" + +primcorec obsf_resource :: "('in, 'out) resource \ ('in, 'out) resource_obsf" where + "run_resource (obsf_resource res) = (\x. + map_spmf (map_prod id obsf_resource) + (map_spmf (map_prod id (\resF. case resF of OK res' \ res' | Fault \ fail_resource)) + (TRY map_spmf (map_prod OK OK) (run_resource res x) ELSE return_spmf (Fault, Fault))))" + +lemma obsf_resource_sel: + "run_resource (obsf_resource res) x = + map_spmf (map_prod id (\resF. obsf_resource (case resF of OK res' \ res' | Fault \ fail_resource))) + (TRY map_spmf (map_prod OK OK) (run_resource res x) ELSE return_spmf (Fault, Fault))" + by(simp add: spmf.map_comp prod.map_comp o_def id_def) + +declare obsf_resource.simps [simp del] + +lemma obsf_resource_exception [simp]: "obsf_resource fail_resource = const_resource Fault" + by coinduction(simp add: rel_fun_def obsf_resource_sel) + +lemma obsf_resource_sel2 [simp]: + "run_resource (obsf_resource res) x = + try_spmf (map_spmf (map_prod OK obsf_resource) (run_resource res x)) (return_spmf (Fault, const_resource Fault))" + by(simp add: map_try_spmf spmf.map_comp o_def prod.map_comp obsf_resource_sel) + +lemma lossless_obsf_resource [simp]: "lossless_resource \ (obsf_resource res)" + by(coinduction arbitrary: res) auto + +lemma WT_obsf_resource [WT_intro, simp]: "exception_\ \ \res obsf_resource res \" if "\ \res res \" + using that by(coinduction arbitrary: res)(auto dest: WT_resourceD split: if_split_asm) + + +type_synonym ('a, 'b) distinguisher_obsf = "(bool, 'a, 'b exception) gpv" + +translations + (type) "('a, 'b) distinguisher_obsf" <= (type) "(bool, 'a, 'b exception) gpv" + +abbreviation connect_obsf :: "('a, 'b) distinguisher_obsf \ ('a, 'b) resource_obsf \ bool spmf" where + "connect_obsf == connect" + +definition obsf_distinguisher :: "('a, 'b) distinguisher \ ('a, 'b) distinguisher_obsf" where + "obsf_distinguisher \ = map_gpv' (\x. x = Some True) id option_of_exception (gpv_stop \)" + +lemma WT_obsf_distinguisher [WT_intro]: + "exception_\ \ \g obsf_distinguisher \ \" if [WT_intro]: "\ \g \ \" + unfolding obsf_distinguisher_def by(rule WT_intro|simp)+ + +lemma interaction_bounded_by_obsf_distinguisher [interaction_bound]: + "interaction_bounded_by consider (obsf_distinguisher \) bound" + if [interaction_bound]: "interaction_bounded_by consider \ bound" + unfolding obsf_distinguisher_def by(rule interaction_bound|simp)+ + +lemma plossless_obsf_distinguisher [simp]: + "plossless_gpv (exception_\ \) (obsf_distinguisher \)" + if "plossless_gpv \ \" "\ \g \ \" + using that unfolding obsf_distinguisher_def by(simp) + + +type_synonym ('a, 'b, 'c, 'd) converter_obsf = "('a, 'b exception, 'c, 'd exception) converter" + +translations + (type) "('a, 'b, 'c, 'd) converter_obsf" <= (type) "('a, 'b exception, 'c, 'd exception) converter" + +primcorec obsf_converter :: "('a, 'b, 'c, 'd) converter \ ('a, 'b, 'c, 'd) converter_obsf" where + "run_converter (obsf_converter conv) = (\x. + map_gpv (map_prod id obsf_converter) id + (map_gpv (\convF. case convF of Fault \ (Fault, fail_converter) | OK (a, conv') \ (OK a, conv')) id + (try_gpv (map_gpv' exception_of_option id option_of_exception (gpv_stop (run_converter conv x))) (Done Fault))))" + +lemma obsf_converter_exception [simp]: "obsf_converter fail_converter = const_converter Fault" + by(coinduction)(simp add: rel_fun_def) + +lemma obsf_converter_sel [simp]: + "run_converter (obsf_converter conv) x = + TRY map_gpv' (\y. case y of None \ (Fault, const_converter Fault) | Some(x, conv') \ (OK x, obsf_converter conv')) id option_of_exception + (gpv_stop (run_converter conv x)) + ELSE Done (Fault, const_converter Fault)" + by(simp add: map_try_gpv) + (simp add: map_gpv_conv_map_gpv' map_gpv'_comp o_def option.case_distrib[where h="map_prod _ _"] split_def id_def cong del: option.case_cong) + +declare obsf_converter.sel [simp del] + +lemma exec_gpv_obsf_resource: + defines "exec_gpv1 \ exec_gpv" + and "exec_gpv2 \ exec_gpv" + shows + "exec_gpv1 run_resource (map_gpv' id id option_of_exception (gpv_stop gpv)) (obsf_resource res) \ {(Some x, y)|x y. True} = + map_spmf (map_prod Some obsf_resource) (exec_gpv2 run_resource gpv res)" + (is "?lhs = ?rhs") +proof(rule spmf.leq_antisym) + show "ord_spmf (=) ?lhs ?rhs" unfolding exec_gpv1_def + proof(induction arbitrary: gpv res rule: exec_gpv_fixp_induct_strong) + case adm show ?case by simp + case bottom show ?case by simp + case (step exec_gpv') + show ?case unfolding exec_gpv2_def + apply(subst exec_gpv.simps) + apply(clarsimp simp add: map_bind_spmf bind_map_spmf restrict_bind_spmf o_def try_spmf_def intro!: ord_spmf_bind_reflI split!: generat.split) + apply(clarsimp simp add: bind_map_pmf bind_spmf_def bind_assoc_pmf bind_return_pmf spmf.leq_trans[OF restrict_spmf_mono, OF step.hyps] id_def step.IH[simplified, simplified id_def exec_gpv2_def] intro!: rel_pmf_bind_reflI split!: option.split) + done + qed + + show "ord_spmf (=) ?rhs ?lhs" unfolding exec_gpv2_def + proof(induction arbitrary: gpv res rule: exec_gpv_fixp_induct) + case adm show ?case by simp + case bottom show ?case by simp + case (step exec_gpv') + show ?case unfolding exec_gpv1_def + apply(subst exec_gpv.simps) + apply(clarsimp simp add: bind_map_spmf map_bind_spmf restrict_bind_spmf o_def try_spmf_def intro!: ord_spmf_bind_reflI split!: generat.split) + apply(clarsimp simp add: bind_spmf_def bind_assoc_pmf bind_map_pmf map_bind_pmf bind_return_pmf id_def step.IH[simplified, simplified id_def exec_gpv1_def] intro!: rel_pmf_bind_reflI split!: option.split) + done + qed +qed + +lemma obsf_attach: + assumes pfinite: "pfinite_converter \ \' conv" + and WT: "\, \' \\<^sub>C conv \" + and WT_resource: "\' \res res \" + shows "outs_\ \ \\<^sub>R attach (obsf_converter conv) (obsf_resource res) \ obsf_resource (attach conv res)" + using assms +proof(coinduction arbitrary: conv res) + case (eq_resource_on out conv res) + then show ?case (is "rel_spmf ?X ?lhs ?rhs") + proof - + have "?lhs = map_spmf (\((b, conv'), res'). (b, conv' \ res')) + (exec_gpv run_resource + (TRY map_gpv' (case_option (Fault, const_converter Fault) (\(x, conv'). (OK x, obsf_converter conv'))) id option_of_exception (gpv_stop (run_converter conv out)) ELSE Done (Fault, const_converter Fault)) + (obsf_resource res))" + (is "_ = map_spmf ?attach (exec_gpv _ (TRY ?gpv ELSE _) _)") by(clarsimp) + also have "\ = TRY map_spmf ?attach (exec_gpv run_resource ?gpv (obsf_resource res)) ELSE return_spmf (Fault, const_resource Fault)" + by(rule run_lossless_resource.exec_gpv_try_gpv[where \="exception_\ \'"]) + (use eq_resource_on in \auto intro!: WT_gpv_map_gpv' WT_gpv_stop pfinite_gpv_stop[THEN iffD2] dest: WT_converterD pfinite_converterD lossless_resourceD\) + also have "\ = TRY map_spmf (\(rc, res'). case rc of None \ (Fault, const_resource Fault) | Some (x, conv') \ (OK x, obsf_converter conv' \ res')) + ((exec_gpv run_resource (map_gpv' id id option_of_exception (gpv_stop (run_converter conv out))) (obsf_resource res)) \ {(Some x, y)|x y. True}) + ELSE return_spmf (Fault, const_resource Fault)" (is "_ = TRY map_spmf ?f _ ELSE ?z") + by(subst map_gpv'_id12)(clarsimp simp add: map_gpv'_map_gpv_swap exec_gpv_map_gpv_id try_spmf_def restrict_spmf_def bind_map_pmf intro!: bind_pmf_cong[OF refl] split: option.split) + also have "\ = TRY map_spmf ?f (map_spmf (map_prod Some obsf_resource) (exec_gpv run_resource (run_converter conv out) res)) ELSE ?z" + by(simp only: exec_gpv_obsf_resource) + also have "rel_spmf ?X \ ?rhs" using eq_resource_on + by(auto simp add: spmf.map_comp o_def spmf_rel_map intro!: rel_spmf_try_spmf rel_spmf_reflI) + (auto intro!: exI dest: run_resource.in_set_spmf_exec_gpv_into_results_gpv WT_converterD pfinite_converterD run_resource.exec_gpv_invariant) + finally show ?case . + qed +qed + + +lemma colossless_obsf_converter [simp]: + "colossless_converter (exception_\ \) \' (obsf_converter conv)" + by(coinduction arbitrary: conv)(auto split: option.split_asm) + + +lemma WT_obsf_converter [WT_intro]: + "exception_\ \, exception_\ \' \\<^sub>C obsf_converter conv \" if "\, \' \\<^sub>C conv \" + using that + by(coinduction arbitrary: conv)(auto 4 3 dest: WT_converterD results_gpv_stop_SomeD split!: option.splits intro!: WT_intro) + +lemma inline1_gpv_stop_obsf_converter: + defines "inline1a \ inline1" + and "inline1b \ inline1" + shows "bind_spmf (inline1a run_converter (map_gpv' id id option_of_exception (gpv_stop gpv)) (obsf_converter conv)) + (\xy. case xy of Inl (None, conv') \ return_pmf None | Inl (Some x, conv') \ return_spmf (Inl (x, conv')) | Inr y \ return_spmf (Inr y)) = + map_spmf (map_sum (apsnd obsf_converter) + (apsnd (map_prod (\rpv input. case input of Fault \ Done (Fault, const_converter Fault) | OK input' \ + map_gpv' (\res. case res of None \ (Fault, const_converter Fault) | Some (x, conv') \ (OK x, obsf_converter conv')) id option_of_exception (try_gpv (gpv_stop (rpv input')) (Done None))) + (\rpv input. case input of Fault \ Done None | OK input' \ map_gpv' id id option_of_exception (gpv_stop (rpv input')))))) + (inline1b run_converter gpv conv)" + (is "?lhs = ?rhs") +proof(rule spmf.leq_antisym) + show "ord_spmf (=) ?lhs ?rhs" unfolding inline1a_def + proof(induction arbitrary: gpv conv rule: inline1_fixp_induct_strong) + case adm show ?case by simp + case bottom show ?case by simp + case (step inline1') + show ?case unfolding inline1b_def + apply(subst inline1_unfold) + apply(clarsimp simp add: map_spmf_bind_spmf bind_map_spmf spmf.map_comp o_def generat.map_comp intro!: ord_spmf_bind_reflI split!: generat.split) + apply(clarsimp simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_map_pmf bind_return_pmf intro!: rel_pmf_bind_reflI split!: option.split) + subgoal unfolding bind_spmf_def[symmetric] + by(rule ord_spmf_bindI[OF step.hyps, THEN spmf.leq_trans]) + (auto split!: option.split intro!: ord_spmf_bindI[OF step.hyps, THEN spmf.leq_trans] ord_spmf_reflI) + subgoal unfolding bind_spmf_def[symmetric] + by(clarsimp simp add: in_set_spmf[symmetric] inline1b_def split!: generat.split intro!: step.IH[THEN spmf.leq_trans]) + (auto simp add: fun_eq_iff map'_try_gpv split: exception.split) + done + qed + + show "ord_spmf (=) ?rhs ?lhs" unfolding inline1b_def + proof(induction arbitrary: gpv conv rule: inline1_fixp_induct_strong) + case adm show ?case by simp + case bottom show ?case by simp + case (step inline1') + show ?case unfolding inline1a_def + apply(subst inline1_unfold) + apply(clarsimp simp add: map_spmf_bind_spmf bind_map_spmf spmf.map_comp o_def generat.map_comp intro!: ord_spmf_bind_reflI split!: generat.split) + apply(clarsimp simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_map_pmf bind_return_pmf intro!: rel_pmf_bind_reflI split!: option.split) + apply(clarsimp simp add: bind_spmf_def[symmetric] in_set_spmf[symmetric] inline1a_def id_def[symmetric] split!: generat.split intro!: step.IH[THEN spmf.leq_trans]) + apply(auto simp add: fun_eq_iff map'_try_gpv split: exception.split) + done + qed +qed + +lemma inline_gpv_stop_obsf_converter: + "bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop gpv)) (obsf_converter conv)) (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv')) = + bind_gpv (map_gpv' id id option_of_exception (gpv_stop (inline run_converter gpv conv))) (\x. case x of None \ Fail | Some (x', conv) \ Done (Some x', obsf_converter conv))" +proof(coinduction arbitrary: gpv conv rule: gpv_coinduct_bind) + case (Eq_gpv gpv conv) + show "?case TYPE('c \ ('b, 'c, 'd, 'e) converter) TYPE('c \ ('b, 'c, 'd, 'e) converter)" (is "rel_spmf ?X ?lhs ?rhs") + proof - + have "?lhs = map_spmf (\xyz. case xyz of Inl (x, conv) \ Pure (Some x, conv) | Inr (out, rpv, rpv') \ IO out (\input. bind_gpv (bind_gpv (rpv input) (\(x, y). inline run_converter (rpv' x) y)) (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv')))) + (bind_spmf (inline1 run_converter (map_gpv' id id option_of_exception (gpv_stop gpv)) (obsf_converter conv)) + (\xy. case xy of Inl (None, conv') \ return_pmf None | Inl (Some x, conv') \ return_spmf (Inl (x, conv')) | Inr y \ return_spmf (Inr y)))" + (is "_ = map_spmf ?f _") + by(auto simp del: bind_gpv_sel' simp add: bind_gpv.sel map_bind_spmf inline_sel bind_map_spmf o_def intro!: bind_spmf_cong[OF refl] split!: sum.split option.split) + also have "\ = map_spmf ?f (map_spmf (map_sum (apsnd obsf_converter) (apsnd (map_prod (\rpv. case_exception (Done (Fault, const_converter Fault)) + (\input'. map_gpv' (case_option (Fault, const_converter Fault) (\(x, conv'). (OK x, obsf_converter conv'))) id option_of_exception (TRY gpv_stop (rpv input') ELSE Done None))) + (\rpv. case_exception (Done None) (\input'. map_gpv' id id option_of_exception (gpv_stop (rpv input'))))))) + (inline1 run_converter gpv conv))" + by(simp only: inline1_gpv_stop_obsf_converter) + also have "\ = bind_spmf (inline1 run_converter gpv conv) (\y. return_spmf (?f (map_sum (apsnd obsf_converter) + (apsnd (map_prod (\rpv. case_exception (Done (Fault, const_converter Fault)) (\input'. map_gpv' (case_option (Fault, const_converter Fault) (\(x, conv'). (OK x, obsf_converter conv'))) id option_of_exception (TRY gpv_stop (rpv input') ELSE Done None))) + (\rpv. case_exception (Done None) (\input'. map_gpv' id id option_of_exception (gpv_stop (rpv input')))))) + y)))" + by(simp add: map_spmf_conv_bind_spmf) + also have "rel_spmf ?X \ (bind_spmf (inline1 run_converter gpv conv) + (\x. map_spmf (map_generat id id ((\) (case_sum id (\r. bind_gpv r (case_option Fail (\(x', conv). Done (Some x', obsf_converter conv))))))) + (case map_generat id id (map_fun option_of_exception (map_gpv' id id option_of_exception)) + (map_generat Some id (\rpv. case_option (Done None) (\input'. gpv_stop (rpv input'))) + (case x of Inl x \ Pure x + | Inr (out, oracle, rpv) \ IO out (\input. bind_gpv (oracle input) (\(x, y). inline run_converter (rpv x) y)))) of + Pure x \ + map_spmf (map_generat id id ((\) Inl)) (the_gpv (case x of None \ Fail | Some (x', conv) \ Done (Some x', obsf_converter conv))) + | IO out c \ return_spmf (IO out (\input. Inr (c input))))))" + (is "rel_spmf _ _ ?rhs2" is "rel_spmf _ (bind_spmf _ ?L) (bind_spmf _ ?R)") + proof(rule rel_spmf_bind_reflI) + fix x :: "'a \ ('b, 'c, 'd, 'e) converter + 'd \ ('c \ ('b, 'c, 'd, 'e) converter, 'd, 'e) rpv \ ('a, 'b, 'c) rpv" + assume x: "x \ set_spmf (inline1 run_converter gpv conv)" + consider (Inl) a conv' where "x = Inl (a, conv')" | (Inr) out rpv rpv' where "x = Inr (out, rpv, rpv')" by(cases x) auto + then show "rel_spmf ?X (?L x) (?R x)" + proof cases + case Inr + have "\(gpv2 :: ('c \ ('b, 'c, 'd, 'e) converter, 'd, 'e exception) gpv) (gpv2' :: ('c \ ('b, 'c, 'd, 'e) converter, 'd, 'e exception) gpv) f f'. + bind_gpv (map_gpv' (case_option (Fault, const_converter Fault) (\p. (OK (fst p), obsf_converter (snd p)))) id option_of_exception (TRY gpv_stop (rpv input) ELSE Done None)) + (\x. case fst x of Fault \ Fail | OK xa \ bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (rpv' xa))) (snd x)) (\p. case fst p of None \ Fail | Some x' \ Done (Some x', snd p))) = + bind_gpv gpv2 f \ + bind_gpv (map_gpv' id id option_of_exception (gpv_stop (rpv input))) (case_option Fail (\x. bind_gpv (map_gpv' id id option_of_exception (gpv_stop (inline run_converter (rpv' (fst x)) (snd x)))) (case_option Fail (\p. Done (Some (fst p), obsf_converter (snd p)))))) = + bind_gpv gpv2' f' \ + rel_gpv (\x y. \gpv conv. f x = bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop gpv)) (obsf_converter conv)) (\p. case fst p of None \ Fail | Some x' \ Done (Some x', snd p)) \ + f' y = bind_gpv (map_gpv' id id option_of_exception (gpv_stop (inline run_converter gpv conv))) (case_option Fail (\p. Done (Some (fst p), obsf_converter (snd p))))) + (=) gpv2 gpv2'" + (is "\gpv2 gpv2' f f'. ?lhs1 input = _ \ ?rhs1 input = _ \ rel_gpv (?X f f') _ _ _") for input + proof(intro exI conjI) + let ?gpv2 = "bind_gpv (map_gpv' id id option_of_exception (TRY gpv_stop (rpv input) ELSE Done None)) (\x. case x of None \ Fail | Some x \ Done x)" + let ?gpv2' = "bind_gpv (map_gpv' id id option_of_exception (gpv_stop (rpv input))) (\x. case x of None \ Fail | Some x \ Done x)" + let ?f = "\x. bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (rpv' (fst x)))) (obsf_converter (snd x))) (\p. case fst p of None \ Fail | Some x' \ Done (Some x', snd p))" + let ?f' = "\x. bind_gpv (map_gpv' id id option_of_exception (gpv_stop (inline run_converter (rpv' (fst x)) (snd x)))) (case_option Fail (\p. Done (Some (fst p), obsf_converter (snd p))))" + show "?lhs1 input = bind_gpv ?gpv2 ?f" + by(subst map_gpv'_id12[THEN trans, OF map_gpv'_map_gpv_swap]) + (auto simp add: bind_map_gpv o_def bind_gpv_assoc intro!: bind_gpv_cong split!: option.split) + show "?rhs1 input = bind_gpv ?gpv2' ?f'" + by(auto simp add: bind_gpv_assoc id_def[symmetric] intro!: bind_gpv_cong split!: option.split) + show "rel_gpv (?X ?f ?f') (=) ?gpv2 ?gpv2'" using Inr x + by(auto simp add: map'_try_gpv id_def[symmetric] bind_try_Done_Fail intro: gpv.rel_refl_strong) + qed + then show ?thesis using Inr + by(clarsimp split!: sum.split exception.split simp add: rel_fun_def bind_gpv_assoc split_def map_gpv'_bind_gpv exception.case_distrib[where h="\x. bind_gpv (inline _ x _) _"] option.case_distrib[where h="\x. bind_gpv (map_gpv' _ _ _ x) _"] cong: exception.case_cong option.case_cong) + qed simp + qed + moreover have "?rhs2 = ?rhs" + by(simp del: bind_gpv_sel' add: bind_gpv.sel map_bind_spmf inline_sel bind_map_spmf o_def) + ultimately show ?thesis by(simp only:) + qed +qed + +lemma obsf_comp_converter: + assumes WT: "\, \' \\<^sub>C conv1 \" "\', \'' \\<^sub>C conv2 \" + and pfinite1: "pfinite_converter \ \' conv1" + shows "exception_\ \, exception_\ \'' \\<^sub>C obsf_converter (comp_converter conv1 conv2) \ comp_converter (obsf_converter conv1) (obsf_converter conv2)" + using WT pfinite1 supply eq_\_gpv_map_gpv[simp del] +proof(coinduction arbitrary: conv1 conv2) + case eq_\_converter + show ?case (is "eq_\_gpv ?X _ ?lhs ?rhs") + proof - + have "eq_\_gpv (=) (exception_\ \'') ?rhs (TRY map_gpv (\((b, conv1'), conv2'). (b, conv1' \ conv2')) id + (inline run_converter + (map_gpv' + (case_option (Fault, const_converter Fault) + (\(x, conv'). (OK x, obsf_converter conv'))) + id option_of_exception (gpv_stop (run_converter conv1 q))) + (obsf_converter conv2)) ELSE Done (Fault, const_converter Fault))" + (is "eq_\_gpv _ _ _ ?rhs2" is "eq_\_gpv _ _ _ (try_gpv (map_gpv ?f _ ?inline) ?else)") + using eq_\_converter + apply simp + apply(rule run_colossless_converter.inline_try_gpv[where \="exception_\ \'"]) + apply(auto intro!: WT_intro pfinite_gpv_stop[THEN iffD2] dest: WT_converterD pfinite_converterD colossless_converterD) + done + term "bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (run_converter conv1 q))) (obsf_converter conv2)) + (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv'))" + + also have "?rhs2 = try_gpv (map_gpv ?f id + (map_gpv (\(xo, conv'). case xo of None \ ((Fault, const_converter Fault), conv') | Some (x, conv) \ ((OK x, obsf_converter conv), conv')) id + (bind_gpv (inline run_converter (map_gpv' id id option_of_exception (gpv_stop (run_converter conv1 q))) (obsf_converter conv2)) + (\(x, conv'). case x of None \ Fail | Some x' \ Done (x, conv'))))) + ?else" + apply(simp add: map_gpv_bind_gpv gpv.map_id) + apply(subst try_gpv_bind_gpv) + apply(simp add: split_def option.case_distrib[where h="map_gpv _ _"] option.case_distrib[where h="fst"] option.case_distrib[where h="\x. try_gpv x _"] cong del: option.case_cong) + apply(subst option.case_distrib[where h=Done, symmetric, abs_def])+ + apply(fold map_gpv_conv_bind) + apply(simp add: map_try_gpv gpv.map_comp o_def) + apply(rule try_gpv_cong) + apply(subst map_gpv'_id12) + apply(subst map_gpv'_map_gpv_swap) + apply(simp add: inline_map_gpv gpv.map_comp id_def[symmetric]) + apply(rule gpv.map_cong[OF refl]) + apply(auto split: option.split) + done + also have "\ = try_gpv (map_gpv ?f id + (map_gpv (\(xo, conv'). case xo of None \ ((Fault, const_converter Fault), conv') | Some (x, conv) \ ((OK x, obsf_converter conv), conv')) id +(bind_gpv + (map_gpv' id id option_of_exception + (gpv_stop (inline run_converter (run_converter conv1 q) conv2))) + (case_option Fail + (\(x', conv). + Done + (Some x', + obsf_converter + conv)))))) ?else" + by(simp only: inline_gpv_stop_obsf_converter) + also have "eq_\_gpv ?X (exception_\ \'') ?lhs \" using eq_\_converter + apply simp + apply(simp add: map_gpv_bind_gpv gpv.map_id) + apply(subst try_gpv_bind_gpv) + apply(simp add: split_def option.case_distrib[where h="map_gpv _ _"] option.case_distrib[where h="fst"] option.case_distrib[where h="\x. try_gpv x _"] cong del: option.case_cong) + apply(subst option.case_distrib[where h=Done, symmetric, abs_def])+ + apply(fold map_gpv_conv_bind) + apply(simp add: map_try_gpv gpv.map_comp o_def) + apply(rule eq_\_gpv_try_gpv_cong) + apply(subst map_gpv'_id12) + apply(subst map_gpv'_map_gpv_swap) + apply(simp add: eq_\_gpv_map_gpv id_def[symmetric]) + apply(subst map_gpv_conv_map_gpv') + apply(subst gpv_stop_map') + apply(subst option.map_id0) + apply(subst map_gpv_conv_map_gpv'[symmetric]) + apply(subst map_gpv'_map_gpv_swap) + apply(simp add: eq_\_gpv_map_gpv id_def[symmetric]) + apply(rule eq_\_gpv_reflI) + apply(clarsimp split!: option.split simp add: eq_onp_def) + apply(erule notE) + apply(rule eq_\_converter_reflI) + apply simp + apply(drule results_gpv_stop_SomeD) + apply(rule conjI) + apply(rule imageI) + apply(drule run_converter.results_gpv_inline) + apply(erule (1) WT_converterD) + apply simp + apply clarsimp + apply(drule (2) WT_converterD_results) + apply simp + apply(rule disjI1) + apply(rule exI conjI refl)+ + apply(drule run_converter.results_gpv_inline) + apply(erule (1) WT_converterD) + apply simp + apply clarsimp + apply(drule (2) WT_converterD_results) + apply simp + apply(drule run_converter.results_gpv_inline) + apply(erule (1) WT_converterD) + apply simp + apply clarsimp + apply(drule (1) pfinite_converterD) + apply blast + apply(rule WT_intro run_converter.WT_gpv_inline_invar|simp)+ + apply(erule (1) WT_converterD) + apply simp + apply(simp add: eq_onp_def) + apply(rule disjI2) + apply(rule eq_\_converter_reflI) + apply simp + done + finally (eq_\_gpv_eq_OO2) show ?thesis . + qed +qed + +lemma resource_of_obsf_oracle_Fault [simp]: + "resource_of_oracle (obsf_oracle oracle) Fault = const_resource Fault" + by(coinduction)(auto simp add: rel_fun_def) + +lemma obsf_resource_of_oracle [simp]: + "obsf_resource (resource_of_oracle oracle s) = resource_of_oracle (obsf_oracle oracle) (OK s)" + by(coinduction arbitrary: s rule: resource.coinduct_strong) + (auto 4 3 simp add: rel_fun_def map_try_spmf spmf_rel_map intro!: rel_spmf_try_spmf rel_spmf_reflI) + +lemma trace_callee_eq_obsf_Fault [simp]: "A \\<^sub>C obsf_oracle callee1(Fault) \ obsf_oracle callee2(Fault)" + by(coinduction rule: trace_callee_eq_coinduct) auto + +lemma obsf_resource_eq_\_cong: "A \\<^sub>R obsf_resource res1 \ obsf_resource res2" if "A \\<^sub>R res1 \ res2" + using that by(coinduction arbitrary: res1 res2)(fastforce intro!: rel_spmf_try_spmf simp add: spmf_rel_map elim!: rel_spmf_mono dest: eq_resource_onD) + +lemma trace_callee_eq_obsf_oracleI: + assumes "trace_callee_eq callee1 callee2 A p q" + shows "trace_callee_eq (obsf_oracle callee1) (obsf_oracle callee2) A (try_spmf (map_spmf OK p) (return_spmf Fault)) (try_spmf (map_spmf OK q) (return_spmf Fault))" + using assms +proof(coinduction arbitrary: p q rule: trace_callee_eq_coinduct_strong) + case (step z p q) + have "?lhs = map_pmf (\x. case x of None \ Some Fault | Some y \ Some (OK y)) (bind_spmf p (\s'. map_spmf fst (callee1 s' z)))" + by(auto simp add: bind_spmf_def try_spmf_def bind_assoc_pmf map_bind_pmf bind_map_pmf bind_return_pmf option.case_distrib[where h="map_pmf _"] option.case_distrib[where h=return_pmf, symmetric, abs_def] map_pmf_def[symmetric] pmf.map_comp o_def intro!: bind_pmf_cong[OF refl] pmf.map_cong[OF refl] split: option.split) + also have "bind_spmf p (\s'. map_spmf fst (callee1 s' z)) = bind_spmf q (\s'. map_spmf fst (callee2 s' z))" + using step(1)[THEN trace_callee_eqD[where xs="[]" and x=z]] step(2) by simp + also have "map_pmf (\x. case x of None \ Some Fault | Some y \ Some (OK y)) \ = ?rhs" + by(auto simp add: bind_spmf_def try_spmf_def bind_assoc_pmf map_bind_pmf bind_map_pmf bind_return_pmf option.case_distrib[where h="map_pmf _"] option.case_distrib[where h=return_pmf, symmetric, abs_def] map_pmf_def[symmetric] pmf.map_comp o_def intro!: bind_pmf_cong[OF refl] pmf.map_cong[OF refl] split: option.split) + finally show ?case . +next + case (sim x s1 s2 ye s1' s2' p q) + have eq1: "bind_spmf (try_spmf (map_spmf OK p) (return_spmf Fault)) (\s. obsf_oracle callee1 s x) = + try_spmf (bind_spmf p (\s. map_spmf (map_prod OK OK) (callee1 s x))) (return_spmf (Fault, Fault))" + by(auto simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_map_pmf bind_return_pmf intro!: bind_pmf_cong[OF refl] split: option.split) + have eq2: "bind_spmf (try_spmf (map_spmf OK q) (return_spmf Fault)) (\s. obsf_oracle callee2 s x) = + try_spmf (bind_spmf q (\s. map_spmf (map_prod OK OK) (callee2 s x))) (return_spmf (Fault, Fault))" + by(auto simp add: bind_spmf_def try_spmf_def bind_assoc_pmf bind_map_pmf bind_return_pmf intro!: bind_pmf_cong[OF refl] split: option.split) + + show ?case + proof(cases ye) + case [simp]: Fault + have "lossless_spmf (bind_spmf p (\s. map_spmf (map_prod OK OK) (callee1 s x))) \ lossless_spmf (bind_spmf q (\s. map_spmf (map_prod OK OK) (callee2 s x)))" + using sim(1)[THEN trace_callee_eqD[where xs="[]" and x=x], THEN arg_cong[where f="lossless_spmf"]] sim(2) by simp + then have "?eq" by(simp add: eq1 eq2)(subst (1 2) cond_spmf_fst_try2, auto) + then show ?thesis .. + next + case [simp]: (OK y) + have eq3: "fst ` set_spmf (bind_spmf p (\s. callee1 s x)) = fst ` set_spmf (bind_spmf q (\s. callee2 s x))" + using trace_callee_eqD[OF sim(1) _ sim(2), where xs="[]", THEN arg_cong[where f="set_spmf"]] + by(auto simp add: bind_UNION image_UN del: equalityI) + show ?thesis + proof(cases "y \ fst ` set_spmf (bind_spmf p (\s. callee1 s x))") + case True + have eq4: "cond_spmf_fst (bind_spmf p (\s. map_spmf (apfst OK) (callee1 s x))) (OK y) = cond_spmf_fst (bind_spmf p (\s. callee1 s x)) y" + "cond_spmf_fst (bind_spmf q (\s. map_spmf (apfst OK) (callee2 s x))) (OK y) = cond_spmf_fst (bind_spmf q (\s. callee2 s x)) y" + by(fold map_bind_spmf[unfolded o_def])(simp_all add: cond_spmf_fst_map_inj) + have ?sim + unfolding eq1 eq2 + apply(subst (1 2) cond_spmf_fst_try1) + apply simp + apply simp + apply(rule exI[where x="cond_spmf_fst (bind_spmf p (\s. map_spmf (apfst OK) (callee1 s x))) ye"]) + apply(rule exI[where x="cond_spmf_fst (bind_spmf q (\s. map_spmf (apfst OK) (callee2 s x))) ye"]) + apply(subst (1 2) try_spmf_lossless) + subgoal using True unfolding eq3 by(auto simp add: bind_UNION image_UN intro!: rev_bexI rev_image_eqI) + subgoal using True by(auto simp add: bind_UNION image_UN intro!: rev_bexI rev_image_eqI) + apply(simp add: map_cond_spmf_fst map_bind_spmf o_def spmf.map_comp map_prod_def split_def) + apply(simp add: eq4) + apply(rule trace_callee_eqI) + subgoal for xs z + using sim(1)[THEN trace_callee_eqD[where xs="(x, y) # xs" and x=z]] sim(2) + apply simp + done + done + then show ?thesis .. + next + case False + with eq3 have "y \ fst ` set_spmf (bind_spmf q (\s. callee2 s x))" by auto + with False have ?eq + apply simp + apply(subst (1 2) cond_spmf_fst_eq_return_None[THEN iffD2]) + apply(auto simp add: bind_UNION split: if_split_asm intro: rev_image_eqI) + done + then show ?thesis .. + qed + qed +qed + +lemma trace_callee_eq'_obsf_resourceI: + assumes " A \\<^sub>C callee1(s) \ callee2(s')" + shows "A \\<^sub>C obsf_oracle callee1(OK s) \ obsf_oracle callee2(OK s')" + using assms[THEN trace_callee_eq_obsf_oracleI] by simp + +lemma trace_eq_obsf_resourceI: + assumes "A \\<^sub>R res1 \ res2" + shows "A \\<^sub>R obsf_resource res1 \ obsf_resource res2" + using assms + apply(subst (2 4) resource_of_oracle_run_resource[symmetric]) + apply(subst (asm) (1 2) resource_of_oracle_run_resource[symmetric]) + apply(drule trace_callee_eq'_obsf_resourceI) + apply simp + apply(simp add: resource_of_oracle_run_resource) + done + +lemma spmf_run_obsf_oracle_obsf_distinguisher [rule_format]: + defines "eg1 \ exec_gpv" and "eg2 \ exec_gpv" shows + "spmf (map_spmf fst (eg1 (obsf_oracle oracle) (obsf_distinguisher gpv) (OK s))) True = + spmf (map_spmf fst (eg2 oracle gpv s)) True" + (is "?lhs = ?rhs") +proof(rule antisym) + show "?lhs \ ?rhs" unfolding eg1_def + proof(induction arbitrary: gpv s rule: exec_gpv_fixp_induct_strong) + case adm show ?case by simp + case bottom show ?case by simp + case (step exec_gpv') + show ?case unfolding eg2_def + apply(subst exec_gpv.simps) + apply(clarsimp simp add: obsf_distinguisher_def bind_map_spmf map_bind_spmf o_def) + apply(subst (1 2) spmf_bind) + apply(rule Bochner_Integration.integral_mono) + apply(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(clarsimp split: generat.split simp add: map_bind_spmf o_def) + apply(simp add: try_spmf_def bind_spmf_pmf_assoc bind_map_pmf) + apply(simp add: bind_spmf_def) + apply(subst (1 2) pmf_bind) + apply(rule Bochner_Integration.integral_mono) + apply(rule measure_pmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(rule measure_pmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(clarsimp split!: option.split simp add: bind_return_pmf) + apply(rule antisym) + apply(rule order_trans) + apply(rule step.hyps[THEN ord_spmf_map_spmfI, THEN ord_spmf_eq_leD]) + apply simp + apply(simp) + apply(rule step.IH[unfolded eg2_def obsf_distinguisher_def]) + done + qed + + show "?rhs \ ?lhs" unfolding eg2_def + proof(induction arbitrary: gpv s rule: exec_gpv_fixp_induct_strong) + case adm show ?case by simp + case bottom show ?case by simp + case (step exec_gpv') + show ?case unfolding eg1_def + apply(subst exec_gpv.simps) + apply(clarsimp simp add: obsf_distinguisher_def bind_map_spmf map_bind_spmf o_def) + apply(subst (1 2) spmf_bind) + apply(rule Bochner_Integration.integral_mono) + apply(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(rule measure_spmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(clarsimp split: generat.split simp add: map_bind_spmf o_def) + apply(simp add: try_spmf_def bind_spmf_pmf_assoc bind_map_pmf) + apply(simp add: bind_spmf_def) + apply(subst (1 2) pmf_bind) + apply(rule Bochner_Integration.integral_mono) + apply(rule measure_pmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(rule measure_pmf.integrable_const_bound[where B=1]; simp add: pmf_le_1) + apply(clarsimp split!: option.split simp add: bind_return_pmf) + apply(rule step.IH[unfolded eg1_def obsf_distinguisher_def]) + done + qed +qed + +lemma spmf_obsf_distinguisher_obsf_resource_True: + "spmf (connect_obsf (obsf_distinguisher \) (obsf_resource res)) True = spmf (connect \ res) True" + unfolding connect_def + apply(subst (2) resource_of_oracle_run_resource[symmetric]) + apply(simp add: exec_gpv_resource_of_oracle spmf.map_comp spmf_run_obsf_oracle_obsf_distinguisher) + done + +lemma advantage_obsf_distinguisher: + "advantage (obsf_distinguisher \) (obsf_resource ideal_resource) (obsf_resource real_resource) = + advantage \ ideal_resource real_resource" + unfolding advantage_def by(simp add: spmf_obsf_distinguisher_obsf_resource_True) + +end diff --git a/thys/Constructive_Cryptography_CM/ROOT b/thys/Constructive_Cryptography_CM/ROOT new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/ROOT @@ -0,0 +1,15 @@ +chapter AFP + +session Constructive_Cryptography_CM (AFP) = Constructive_Cryptography + + options [timeout = 1500] + sessions + "Game_Based_Crypto" + "Sigma_Commit_Crypto" + directories + "Constructions" + "Specifications" + theories + "Constructions/DH_OTP" + document_files + "root.tex" + "root.bib" diff --git a/thys/Constructive_Cryptography_CM/Specifications/Channel.thy b/thys/Constructive_Cryptography_CM/Specifications/Channel.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Specifications/Channel.thy @@ -0,0 +1,183 @@ +theory Channel + imports + "../Fused_Resource" +begin + +section \Channel specification\ + +locale ideal_channel = + fixes + (*TODO: decide to either include or exlude this fixation, + - if it is not included, then the types must be manually stated whenever channel locale is imported (c.f. OTP and DH) + - also, the WT_intro rules for channel will have an unspecified valid_messages*) + (* valid_messages :: "'msg set" and*) + leak :: "'msg \ 'leak" and + editable :: bool +begin + + +subsection \Data-types for Parties, State, Events, Input, and Output\ + +datatype party = Alice | Bob + +type_synonym s_shell = "party set" +datatype 'msg' s_kernel = State_Void | State_Store 'msg' | State_Collect 'msg' | State_Collected +type_synonym 'msg' state = "'msg' s_kernel \ s_shell" + +datatype event = Event_Shell party + +datatype iadv_drop = Inp_Drop +datatype iadv_look = Inp_Look +datatype 'msg' iadv_fedit = Inp_Fedit 'msg' +type_synonym 'msg' iadv = "iadv_drop + iadv_look + 'msg' iadv_fedit" + +datatype 'msg' iusr_alice = Inp_Send 'msg' +datatype iusr_bob = Inp_Recv +type_synonym 'msg' iusr = "'msg' iusr_alice + iusr_bob" + +datatype oadv_drop = Out_Drop +datatype 'leak' oadv_look = Out_Look 'leak' +datatype oadv_fedit = Out_Fedit +type_synonym 'leak' oadv = "oadv_drop + 'leak' oadv_look + oadv_fedit" + +datatype ousr_alice = Out_Send +datatype 'msg' ousr_bob = Out_Recv 'msg' +type_synonym 'msg' ousr = "ousr_alice + 'msg' ousr_bob" + + +subsubsection \Basic lemmas for automated handling of party sets (i.e. @{text s_shell})\ + +lemma Alice_neq_iff [simp]: "Alice \ x \ x = Bob" + by(cases x) simp_all + +lemma neq_Alice_iff [simp]: "x \ Alice \ x = Bob" + by(cases x) simp_all + +lemma Bob_neq_iff [simp]: "Bob \ x \ x = Alice" + by(cases x) simp_all + +lemma neq_Bob_iff [simp]: "x \ Bob \ x = Alice" + by(cases x) simp_all + +lemma Alice_in_iff_nonempty: "Alice \ A \ A \ {}" if "Bob \ A" + using that by(auto)(metis (full_types) party.exhaust) + +lemma Bob_in_iff_nonempty: "Bob \ A \ A \ {}" if "Alice \ A" + using that by(auto)(metis (full_types) party.exhaust) + + +subsection \Defining the event handler\ + +fun poke :: "('msg state, event) handler" + where + "poke (s_kernel, parties) (Event_Shell party) = + (if party \ parties then + return_pmf None + else + return_spmf (s_kernel, insert party parties))" + +lemma poke_alt_def: + "poke = (\(s, ps) e. map_spmf (Pair s) (case e of Event_Shell party \ if party \ ps then return_pmf None else return_spmf (insert party ps)))" + by(simp add: fun_eq_iff split: event.split) + +subsection \Defining the adversary interfaces\ + +fun iface_drop :: "('msg state, iadv_drop, oadv_drop) oracle'" + where + "iface_drop _ _ = return_pmf None" + +fun iface_look :: "('msg state, iadv_look, 'leak oadv_look) oracle'" + where + "iface_look (State_Store msg, parties) _ = + return_spmf (Out_Look (leak msg), State_Store msg, parties)" + | "iface_look _ _ = return_pmf None" + +fun iface_fedit :: "('msg state, 'msg iadv_fedit, oadv_fedit) oracle'" + where + "iface_fedit (State_Store msg, parties) (Inp_Fedit msg') = + (if editable then + return_spmf (Out_Fedit, State_Collect msg', parties) + else + return_spmf (Out_Fedit, State_Collect msg, parties))" + | "iface_fedit _ _ = return_pmf None" + +abbreviation iface_adv :: "('msg state, 'msg iadv, 'leak oadv) oracle'" + where + "iface_adv \ plus_oracle iface_drop (plus_oracle iface_look iface_fedit)" + +lemma in_set_spmf_iface_drop: "ys' \ set_spmf (iface_drop s x) \ False" + by simp + +lemma in_set_spmf_iface_look: "ys' \ set_spmf (iface_look s x) \ + (\msg parties. s = (State_Store msg, parties) \ ys' = (Out_Look (leak msg), State_Store msg, parties))" + by(cases "(s, x)" rule: iface_look.cases) simp_all + +lemma in_set_spmf_iface_fedit: "ys' \ set_spmf (iface_fedit s x) \ + (\msg parties msg'. s = (State_Store msg, parties) \ x = (Inp_Fedit msg') \ + ys' = (if editable then (Out_Fedit, State_Collect msg', parties) else (Out_Fedit, State_Collect msg, parties)))" + by(cases "(s, x)" rule: iface_fedit.cases) simp_all + +subsection \Defining the user interfaces\ + +fun iface_alice :: "('msg state, 'msg iusr_alice, ousr_alice) oracle'" + where + "iface_alice (State_Void, parties) (Inp_Send msg) = + (if Alice \ parties then + return_spmf (Out_Send, State_Store msg, parties) + else + return_pmf None)" + | "iface_alice _ _ = return_pmf None" + +fun iface_bob :: "('msg state, iusr_bob, 'msg ousr_bob) oracle'" + where + "iface_bob (State_Collect msg, parties) _ = + (if Bob \ parties then + return_spmf (Out_Recv msg, State_Collected, parties) + else + return_pmf None)" + | "iface_bob _ _ = return_pmf None" + +abbreviation iface_usr :: "('msg state, 'msg iusr, 'msg ousr) oracle'" + where + "iface_usr \ plus_oracle iface_alice iface_bob" + +lemma in_set_spmf_iface_alice: "ys' \ set_spmf (iface_alice s x) \ + (\parties msg. s = (State_Void, parties) \ x = Inp_Send msg \ Alice \ parties \ ys' = (Out_Send, State_Store msg, parties))" + by(cases "(s, x)" rule: iface_alice.cases) simp_all + +lemma in_set_spmf_iface_bob: "ys' \ set_spmf (iface_bob s x) \ + (\msg parties. s = (State_Collect msg, parties) \ Bob \ parties \ ys' = (Out_Recv msg, State_Collected, parties))" + by(cases "(s, x)" rule: iface_bob.cases) simp_all + +subsection \Defining the Fused Resource\ + +primcorec core :: "('msg state, event, 'msg iadv, 'msg iusr, 'leak oadv, 'msg ousr) core" + where + "cpoke core = poke" + | "cfunc_adv core = iface_adv" + | "cfunc_usr core = iface_usr" + +sublocale fused_resource core "(State_Void, {})" . + + +subsubsection \Lemma showing that the resulting resource is well-typed\ + +lemma WT_core [WT_intro]: + "WT_core (\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (Inp_Fedit ` valid_messages) UNIV)) (\_uniform (Inp_Send ` valid_messages) UNIV \\<^sub>\ (\_uniform UNIV (Out_Recv ` valid_messages))) + (pred_prod (pred_s_kernel (\msg. msg \ valid_messages)) (\_. True)) core" + apply(rule WT_core.intros) + subgoal for s e s' by(cases "(s, e)" rule: poke.cases)(auto split: if_split_asm) + subgoal for s x y s' by(cases "(s, projl (projr x))" rule: iface_look.cases)(auto split: if_split_asm) + subgoal for s x y s' by(cases "(s, projl x)" rule: iface_alice.cases)(auto split: if_split_asm) + done + +lemma WT_fuse [WT_intro]: + assumes [WT_intro]: "WT_rest \_adv_rest \_usr_rest I_rest rest" + shows "((\_full \\<^sub>\ (\_full \\<^sub>\ \_uniform (Inp_Fedit ` valid_messages) UNIV)) \\<^sub>\ \_adv_rest) \\<^sub>\ + ((\_uniform (Inp_Send ` valid_messages) UNIV \\<^sub>\ \_uniform UNIV (Out_Recv ` valid_messages)) \\<^sub>\ \_usr_rest) \res resource rest \" + by(rule WT_intro)+ simp + + +end + +end diff --git a/thys/Constructive_Cryptography_CM/Specifications/Key.thy b/thys/Constructive_Cryptography_CM/Specifications/Key.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/Specifications/Key.thy @@ -0,0 +1,151 @@ +theory Key + imports + "../Fused_Resource" +begin + + +section \Key specification\ + +locale ideal_key = + fixes valid_keys :: "'key set" +begin + +subsection \Data-types for Parties, State, Events, Input, and Output\ + +datatype party = Alice | Bob + +type_synonym s_shell = "party set" +datatype 'key' s_kernel = PState_Store | State_Store 'key' +type_synonym 'key' state = "'key' s_kernel \ s_shell" + +datatype event = Event_Shell party | Event_Kernel + +datatype iadv = Inp_Adversary + +datatype iusr_alice = Inp_Alice +datatype iusr_bob = Inp_Bob +type_synonym iusr = "iusr_alice + iusr_bob" + +datatype oadv = Out_Adversary + +datatype 'key' ousr_alice = Out_Alice 'key' +datatype 'key' ousr_bob = Out_Bob 'key' +type_synonym 'key' ousr = "'key' ousr_alice + 'key' ousr_bob" + + +subsubsection \Basic lemmas for automated handling of party sets (i.e. @{text s_shell})\ + +lemma Alice_neq_iff [simp]: "Alice \ x \ x = Bob" + by(cases x) simp_all + +lemma neq_Alice_iff [simp]: "x \ Alice \ x = Bob" + by(cases x) simp_all + +lemma Bob_neq_iff [simp]: "Bob \ x \ x = Alice" + by(cases x) simp_all + +lemma neq_Bob_iff [simp]: "x \ Bob \ x = Alice" + by(cases x) simp_all + +lemma Alice_in_iff_nonempty: "Alice \ A \ A \ {}" if "Bob \ A" + using that by(auto)(metis (full_types) party.exhaust) + +lemma Bob_in_iff_nonempty: "Bob \ A \ A \ {}" if "Alice \ A" + using that by(auto)(metis (full_types) party.exhaust) + + +subsection \Defining the event handler\ + +fun poke :: "('key state, event) handler" + where + "poke (s_kernel, parties) (Event_Shell party) = + (if party \ parties then + return_pmf None + else + return_spmf (s_kernel, insert party parties))" + | "poke (PState_Store, s_shell) (Event_Kernel) = do { + key \ spmf_of_set valid_keys; + return_spmf (State_Store key, s_shell) } " + | "poke _ _ = return_pmf None" + +lemma in_set_spmf_poke: + "s' \ set_spmf (poke s x) \ + (\s_kernel parties party. s = (s_kernel, parties) \ x = Event_Shell party \ party \ parties \ s' = (s_kernel, insert party parties)) \ + (\s_shell key. s = (PState_Store, s_shell) \ x = Event_Kernel \ key \ valid_keys \ finite valid_keys \ s' = (State_Store key, s_shell))" + by(cases "(s, x)" rule: poke.cases)(auto simp add: set_spmf_of_set) + +lemma foldl_poke_invar: + "\ (s_kernel', parties') \ set_spmf (foldl_spmf poke p events); \(s_kernel, parties)\set_spmf p. set_s_kernel s_kernel \ valid_keys \ + \ set_s_kernel s_kernel' \ valid_keys" + by(induction events arbitrary: parties' rule: rev_induct) + (auto 4 3 simp add: split_def foldl_spmf_append in_set_spmf_poke dest: bspec) + +subsection \Defining the adversary interface\ + +fun iface_adv :: "('key state, iadv, oadv) oracle'" + where + "iface_adv state _ = return_spmf (Out_Adversary, state)" + + +subsection \Defining the user interfaces\ + +context +begin + +private fun iface_usr_func :: "party \ _ \ _ \ 'inp \ ('wrap_key \ 'key state) spmf" + where + "iface_usr_func party wrap (State_Store key, parties) inp = + (if party \ parties then + return_spmf (wrap key, State_Store key, parties) + else + return_pmf None)" + | "iface_usr_func _ _ _ _ = return_pmf None" + +abbreviation iface_alice :: "('key state, iusr_alice, 'key ousr_alice) oracle'" + where + "iface_alice \ iface_usr_func Alice Out_Alice" + +abbreviation iface_bob :: "('key state, iusr_bob, 'key ousr_bob) oracle'" + where + "iface_bob \ iface_usr_func Bob Out_Bob" + +abbreviation iface_usr :: "('key state, iusr, 'key ousr) oracle'" + where + "iface_usr \ plus_oracle iface_alice iface_bob" + +lemma in_set_iface_usr_func [simp]: + "x \ set_spmf (iface_usr_func party wrap state inp) \ + (\key parties. state = (State_Store key, parties) \ party \ parties \ x = (wrap key, State_Store key, parties))" + by(cases "(party, wrap, state, inp)" rule: iface_usr_func.cases) auto + +end + + +subsection \Defining the Fuse Resource\ + +primcorec core :: "('key state, event, iadv, iusr, oadv,'key ousr) core" + where + "cpoke core = poke" + | "cfunc_adv core = iface_adv" + | "cfunc_usr core = iface_usr" + +sublocale fused_resource core "(PState_Store, {})" . + + +subsubsection \Lemma showing that the resulting resource is well-typed\ + +lemma WT_core [WT_intro]: + "WT_core \_full (\_uniform UNIV (Out_Alice ` valid_keys) \\<^sub>\ \_uniform UNIV (Out_Bob ` valid_keys)) + (pred_prod (pred_s_kernel (\key. key \ valid_keys)) (\_. True)) core" + apply (rule WT_core.intros) + subgoal for s e s' by(cases "(s, e)" rule: poke.cases)(auto split: if_split_asm simp add: set_spmf_of_set) + by auto + +lemma WT_fuse [WT_intro]: + assumes [WT_intro]: "WT_rest \_adv_rest \_usr_rest I_rest rest" + shows "(\_full \\<^sub>\ \_adv_rest) \\<^sub>\ ((\_uniform UNIV (Out_Alice ` valid_keys) \\<^sub>\ \_uniform UNIV (Out_Bob ` valid_keys)) \\<^sub>\ \_usr_rest) \res resource rest \" + by(rule WT_intro)+ simp + +end + +end diff --git a/thys/Constructive_Cryptography_CM/State_Isomorphism.thy b/thys/Constructive_Cryptography_CM/State_Isomorphism.thy new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/State_Isomorphism.thy @@ -0,0 +1,89 @@ +theory State_Isomorphism + imports + More_CC +begin + + +section \State Isomorphism\ + +type_synonym + ('a, 'b) state_iso = "('a \ 'b) \ ('b \ 'a)" + +definition + state_iso :: "('a, 'b) state_iso \ bool" + where + "state_iso \ (\(f, g). type_definition f g UNIV)" + +definition + apply_state_iso :: " ('s1, 's2) state_iso \ ('s1, 'i, 'o) oracle' \ ('s2, 'i, 'o) oracle'" + where + "apply_state_iso \ (\(f, g). map_fun g (map_fun id (map_spmf (map_prod id f))))" + +lemma apply_state_iso_id: "apply_state_iso (id, id) = id" + by (auto simp add: apply_state_iso_def map_prod.id spmf.map_id0 map_fun_id) + +lemma apply_state_iso_compose: "apply_state_iso si1 (apply_state_iso si2 oracle) = + apply_state_iso (map_prod (\f. f o (fst si2)) ((o) (snd si2)) si1) oracle" + unfolding apply_state_iso_def + by (auto simp add: split_def id_def o_def map_prod_def map_fun_def map_spmf_conv_bind_spmf) + +lemma apply_wiring_state_iso_assoc: + "apply_wiring wr (apply_state_iso si oracle) = apply_state_iso si (apply_wiring wr oracle)" + unfolding apply_state_iso_def apply_wiring_def + by (auto simp add: split_def id_def o_def map_prod_def map_fun_def map_spmf_conv_bind_spmf) + +lemma + resource_of_oracle_state_iso: + assumes "state_iso fg" + shows "resource_of_oracle (apply_state_iso fg oracle) s = resource_of_oracle oracle (snd fg s)" +proof - + have [simp]: "snd fg (fst fg x) = x" for x + using assms by(simp add: state_iso_def split_beta type_definition.Rep_inverse) + show ?thesis + by(coinduction arbitrary: s) + (auto 4 3 simp add: rel_fun_def spmf_rel_map apply_state_iso_def split_def intro!: rel_spmf_reflI) +qed + + +subsection \Parallel State Isomorphism\ + +definition + parallel_state_iso :: " (('s_core1 \ 's_core2) \ ('s_rest1 \ 's_rest2), + ('s_core1 \ 's_rest1) \ ('s_core2 \ 's_rest2)) state_iso" + where + "parallel_state_iso = + (\((s11, s12), (s21, s22)). ((s11, s21), (s12, s22)), + \((s11, s21), (s12, s22)). ((s11, s12), (s21, s22)))" + +lemma + state_iso_parallel_state_iso [simp]: "state_iso parallel_state_iso" + by (auto simp add: type_definition_def state_iso_def parallel_state_iso_def) + + + +subsection \Trisplit State Isomorphism\ + +definition + iso_trisplit + where + "iso_trisplit = + (\(((s11, s12), s13), (s21, s22), s23). (((s11, s21), s12, s22), s13, s23), + \(((s11, s21), s12, s22), s13, s23). (((s11, s12), s13), (s21, s22), s23))" + +lemma + state_iso_fuse_par [simp]: "state_iso iso_trisplit" + by(simp add: state_iso_def iso_trisplit_def; unfold_locales; simp add: split_def) + + +subsection \Assocl-Swap State Isomorphism\ + +definition + iso_swapar + where + "iso_swapar = (\((sm, s1), s2). (s1, sm, s2), \(s1, sm, s2). ((sm, s1), s2))" + +lemma + state_iso_swapar [simp]: "state_iso iso_swapar" + by(simp add: state_iso_def iso_swapar_def; unfold_locales; simp add: split_def) + +end \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/document/root.bib b/thys/Constructive_Cryptography_CM/document/root.bib new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/document/root.bib @@ -0,0 +1,87 @@ +@InProceedings{Maurer2011, + author = {Ueli Maurer and Renato Renner}, + title = {Abstract Cryptography}, + booktitle = {Innovations in Computer Science ({ICS} 2010), Proceedings}, + year = {2011}, + pages = {1--21}, + publisher = {Tsinghua University Press}, +} + +@InProceedings{Maurer2011a, + author = {Ueli Maurer}, + title = {Constructive Cryptography - {A} New Paradigm for Security Definitions and Proofs}, + booktitle = {Theory of Security and Applications - Joint Workshop ({TOSCA} 2011), Revised Selected Papers}, + year = {2011}, + volume = {6993}, + series = {LNCS}, + pages = {33--56}, + publisher = {Springer} +} + +@InProceedings{Maurer2016, + author = {Ueli Maurer and + Renato Renner}, + title = {From Indifferentiability to Constructive Cryptography (and Back)}, + booktitle = {Theory of Cryptography Conference ({TCC} 2016), Proceedings, Part {I}}, + year = {2016}, + volume = {9985}, + series = {LNCS}, + pages = {3--24}, + publisher = {Springer} +} + +@InProceedings{Lochbihler2016, + author = {Andreas Lochbihler}, + title = {Probabilistic Functions and Cryptographic Oracles in Higher Order Logic}, + booktitle = {European Symposium on Programming ({ESOP} 2016), Proceedings}, + year = {2016}, + volume = {9632}, + series = {LNCS}, + pages = {503--531}, + publisher = {Springer}, +} + +@Article{Lochbihler2017, + author = {Andreas Lochbihler}, + title = {{CryptHOL}}, + journal = {Archive of Formal Proofs}, + year = 2017, + note = {~\url{https://isa-afp.org/entries/CryptHOL.html}, + Formal proof development}, +} + +@Article{Lochbihler2018, + author = {Andreas Lochbihler and S. Reza Sefidgar}, + title = {Constructive Cryptography in {HOL}}, + journal = {Archive of Formal Proofs}, + year = 2018, + note = {~\url{https://isa-afp.org/entries/Constructive_Cryptography.html}, + Formal proof development}, +} + +@InProceedings{Lochbihler2019, + author = {Andreas Lochbihler and S. Reza Sefidgar and David Basin and Ueli Maurer}, + title = {Formalizing Constructive Cryptography using CryptHOL}, + booktitle = {Computer Security Foundations Symposium ({CSF} 2019), Proceedings}, + year = {2019}, + pages = {152--166}, + publisher = {{IEEE}}, +} + +@Article{Basin2020, + author = {David A. Basin and Andreas Lochbihler and S. Reza Sefidgar}, + title = {{CryptHOL}: Game-based Proofs in Higher-order Logic}, + journal = {Journal of Cryptology}, + year = {2020}, + volume = {33}, + number = {2}, + pages = {494--566}, +} + + +@Article{Basin2021, + author = {David A. Basin and Andreas Lochbihler and Ueli Maurer and S. Reza Sefidgar}, + title = {Abstract Modeling of Systems Communication in Constructive Cryptography using {CryptHOL}}, + year = {2021}, + note = {~\url{http://www.andreas-lochbihler.de/pub/basin2021.pdf}, Draft paper} +} \ No newline at end of file diff --git a/thys/Constructive_Cryptography_CM/document/root.tex b/thys/Constructive_Cryptography_CM/document/root.tex new file mode 100644 --- /dev/null +++ b/thys/Constructive_Cryptography_CM/document/root.tex @@ -0,0 +1,58 @@ +\documentclass[11pt,a4paper]{article} +\usepackage{isabelle,isabellesym} +\usepackage{amssymb,amsmath} +\usepackage[english]{babel} +\usepackage[only,bigsqcap]{stmaryrd} +\usepackage{wasysym} +\usepackage{booktabs} + +% this should be the last package used +\usepackage{pdfsetup} + +% urls in roman style, theory text in math-similar italics +\urlstyle{rm} +\isabellestyle{it} + +\begin{document} + +\title{Constructive Cryptography in HOL: the Communication Modeling Aspect} +\author{Andreas Lochbihler and S. Reza Sefidgar} +\maketitle + +\begin{abstract} +Constructive Cryptography (CC)~\cite{Maurer2011, Maurer2011a, Maurer2016} +introduces an abstract approach to composable security statements that allows one +to focus on a particular aspect of security proofs at a time. +Instead of proving the properties of concrete systems, CC studies system +classes, i.e., the shared behavior of similar systems, and their transformations. + +Modeling of systems communication plays a crucial role in composability and +reusability of security statements; yet, this aspect has not been studied +in any of the existing CC results. We extend our previous CC formalization~\cite{Lochbihler2018,Lochbihler2019} +with a new semantic domain called Fused Resource Templates (FRT) that abstracts over +the systems communication patterns in CC proofs. This widens the scope of cryptography proof +formalizations in the {CryptHOL} library~\cite{Lochbihler2017,Lochbihler2016,Basin2020}. + +This formalization is described in~\cite{Basin2021}. +\end{abstract} + +\tableofcontents + +\clearpage + +% sane default for proof documents +\parindent 0pt\parskip 0.5ex + +% generated text of all theories +\input{session} + +% optional bibliography +\bibliographystyle{abbrv} +\bibliography{root} + +\end{document} + +%%% Local Variables: +%%% mode: latex +%%% TeX-master: t +%%% End: diff --git a/thys/ROOTS b/thys/ROOTS --- a/thys/ROOTS +++ b/thys/ROOTS @@ -1,591 +1,592 @@ ADS_Functor AI_Planning_Languages_Semantics AODV AVL-Trees AWN Abortable_Linearizable_Modules Abs_Int_ITP2012 Abstract-Hoare-Logics Abstract-Rewriting Abstract_Completeness Abstract_Soundness Adaptive_State_Counting Affine_Arithmetic Aggregation_Algebras Akra_Bazzi Algebraic_Numbers Algebraic_VCs Allen_Calculus Amicable_Numbers Amortized_Complexity AnselmGod Applicative_Lifting Approximation_Algorithms Architectural_Design_Patterns Aristotles_Assertoric_Syllogistic Arith_Prog_Rel_Primes ArrowImpossibilityGS Attack_Trees Auto2_HOL Auto2_Imperative_HOL AutoFocus-Stream Automated_Stateful_Protocol_Verification Automatic_Refinement AxiomaticCategoryTheory BDD BNF_CC BNF_Operations BTree Banach_Steinhaus Bell_Numbers_Spivey Berlekamp_Zassenhaus Bernoulli Bertrands_Postulate Bicategory BinarySearchTree Binding_Syntax_Theory Binomial-Heaps Binomial-Queues BirdKMP Blue_Eyes Bondy Boolean_Expression_Checkers Bounded_Deducibility_Security Buchi_Complementation Budan_Fourier Buffons_Needle Buildings BytecodeLogicJmlTypes C2KA_DistributedSystems CAVA_Automata CAVA_LTL_Modelchecker CCS CISC-Kernel CRDT CYK CakeML CakeML_Codegen Call_Arity Card_Equiv_Relations Card_Multisets Card_Number_Partitions Card_Partitions Cartan_FP Case_Labeling Catalan_Numbers Category Category2 Category3 Cauchy Cayley_Hamilton Certification_Monads Chandy_Lamport Chord_Segments Circus Clean ClockSynchInst Closest_Pair_Points CofGroups Coinductive Coinductive_Languages Collections Comparison_Sort_Lower_Bound Compiling-Exceptions-Correctly Complete_Non_Orders Completeness Complex_Geometry Complx ComponentDependencies ConcurrentGC ConcurrentIMP Concurrent_Ref_Alg Concurrent_Revisions Consensus_Refined Constructive_Cryptography +Constructive_Cryptography_CM Constructor_Funs Containers CoreC++ Core_DOM Core_SC_DOM Count_Complex_Roots CryptHOL CryptoBasedCompositionalProperties CSP_RefTK DFS_Framework DPT-SAT-Solver DataRefinementIBP Datatype_Order_Generator Decl_Sem_Fun_PL Decreasing-Diagrams Decreasing-Diagrams-II Deep_Learning Delta_System_Lemma Density_Compiler Dependent_SIFUM_Refinement Dependent_SIFUM_Type_Systems Depth-First-Search Derangements Deriving Descartes_Sign_Rule Dict_Construction Differential_Dynamic_Logic Differential_Game_Logic Dijkstra_Shortest_Path Diophantine_Eqns_Lin_Hom Dirichlet_L Dirichlet_Series DiscretePricing Discrete_Summation DiskPaxos DOM_Components DynamicArchitectures Dynamic_Tables E_Transcendental Echelon_Form EdmondsKarp_Maxflow Efficient-Mergesort Elliptic_Curves_Group_Law Encodability_Process_Calculi Epistemic_Logic Ergodic_Theory Error_Function Euler_MacLaurin Euler_Partition Example-Submission Extended_Finite_State_Machine_Inference Extended_Finite_State_Machines FFT FLP FOL-Fitting FOL_Harrison FOL_Seq_Calc1 Factored_Transition_System_Bounding Falling_Factorial_Sum Farkas FeatherweightJava Featherweight_OCL Fermat3_4 FileRefinement FinFun Finger-Trees Finite-Map-Extras Finite_Automata_HF First_Order_Terms First_Welfare_Theorem Fishburn_Impossibility Fisher_Yates Flow_Networks Floyd_Warshall Flyspeck-Tame FocusStreamsCaseStudies Forcing Formal_Puiseux_Series Formal_SSA Formula_Derivatives Fourier Free-Boolean-Algebra Free-Groups FunWithFunctions FunWithTilings Functional-Automata Functional_Ordered_Resolution_Prover Furstenberg_Topology GPU_Kernel_PL Gabow_SCC Game_Based_Crypto Gauss-Jordan-Elim-Fun Gauss_Jordan Gauss_Sums Gaussian_Integers GenClock General-Triangle Generalized_Counting_Sort Generic_Deriving Generic_Join GewirthPGCProof Girth_Chromatic GoedelGod Goedel_HFSet_Semantic Goedel_HFSet_Semanticless Goedel_Incompleteness Goodstein_Lambda GraphMarkingIBP Graph_Saturation Graph_Theory Green Groebner_Bases Groebner_Macaulay Gromov_Hyperbolicity Group-Ring-Module HOL-CSP HOLCF-Prelude HRB-Slicing Heard_Of Hello_World HereditarilyFinite Hermite Hermite_Lindemann Hidden_Markov_Models Higher_Order_Terms Hoare_Time Hood_Melville_Queue HotelKeyCards Huffman Hybrid_Logic Hybrid_Multi_Lane_Spatial_Logic Hybrid_Systems_VCs HyperCTL IEEE_Floating_Point IMAP-CRDT IMO2019 IMP2 IMP2_Binary_Heap IP_Addresses Imperative_Insertion_Sort Impossible_Geometry Incompleteness Incredible_Proof_Machine Inductive_Confidentiality Inductive_Inference InfPathElimination InformationFlowSlicing InformationFlowSlicing_Inter Integration Interpreter_Optimizations Interval_Arithmetic_Word32 Iptables_Semantics Irrational_Series_Erdos_Straus Irrationality_J_Hancl Isabelle_C Isabelle_Marries_Dirac Isabelle_Meta_Model IsaGeoCoq Jacobson_Basic_Algebra Jinja JinjaDCI JinjaThreads JiveDataStoreModel Jordan_Hoelder Jordan_Normal_Form KAD KAT_and_DRA KBPs KD_Tree Key_Agreement_Strong_Adversaries Kleene_Algebra Knuth_Bendix_Order Knot_Theory Knuth_Bendix_Order Knuth_Morris_Pratt Koenigsberg_Friendship Kruskal Kuratowski_Closure_Complement LLL_Basis_Reduction LLL_Factorization LOFT LTL LTL_Master_Theorem LTL_Normal_Form LTL_to_DRA LTL_to_GBA Lam-ml-Normalization LambdaAuth LambdaMu Lambda_Free_EPO Lambda_Free_KBOs Lambda_Free_RPOs Lambert_W Landau_Symbols Laplace_Transform Latin_Square LatticeProperties Launchbury Laws_of_Large_Numbers Lazy-Lists-II Lazy_Case Lehmer Lifting_Definition_Option LightweightJava LinearQuantifierElim Linear_Inequalities Linear_Programming Linear_Recurrences Liouville_Numbers List-Index List-Infinite List_Interleaving List_Inversions List_Update LocalLexing Localization_Ring Locally-Nameless-Sigma Lowe_Ontological_Argument Lower_Semicontinuous Lp Lucas_Theorem MFMC_Countable MFODL_Monitor_Optimized MFOTL_Monitor MSO_Regex_Equivalence Markov_Models Marriage Mason_Stothers Matrices_for_ODEs Matrix Matrix_Tensor Matroids Max-Card-Matching Median_Of_Medians_Selection Menger Mereology Mersenne_Primes MiniML Minimal_SSA Minkowskis_Theorem Minsky_Machines Modal_Logics_for_NTS Modular_arithmetic_LLL_and_HNF_algorithms Modular_Assembly_Kit_Security Monad_Memo_DP Monad_Normalisation MonoBoolTranAlgebra MonoidalCategory Monomorphic_Monad MuchAdoAboutTwo Multi_Party_Computation Multirelations Myhill-Nerode Name_Carrying_Type_Inference Nash_Williams Nat-Interval-Logic Native_Word Nested_Multisets_Ordinals Network_Security_Policy_Verification Neumann_Morgenstern_Utility No_FTL_observers Nominal2 Noninterference_CSP Noninterference_Concurrent_Composition Noninterference_Generic_Unwinding Noninterference_Inductive_Unwinding Noninterference_Ipurge_Unwinding Noninterference_Sequential_Composition NormByEval Nullstellensatz Octonions OpSets Open_Induction Optics Optimal_BST Orbit_Stabiliser Order_Lattice_Props Ordered_Resolution_Prover Ordinal Ordinal_Partitions Ordinals_and_Cardinals Ordinary_Differential_Equations PAC_Checker PCF PLM POPLmark-deBruijn PSemigroupsConvolution Pairing_Heap Paraconsistency Parity_Game Partial_Function_MR Partial_Order_Reduction Password_Authentication_Protocol Pell Perfect-Number-Thm Perron_Frobenius Physical_Quantities Pi_Calculus Pi_Transcendental Planarity_Certificates Poincare_Bendixson Poincare_Disc Polynomial_Factorization Polynomial_Interpolation Polynomials Pop_Refinement Posix-Lexing Possibilistic_Noninterference Power_Sum_Polynomials Pratt_Certificate Presburger-Automata Prim_Dijkstra_Simple Prime_Distribution_Elementary Prime_Harmonic_Series Prime_Number_Theorem Priority_Queue_Braun Priority_Search_Trees Probabilistic_Noninterference Probabilistic_Prime_Tests Probabilistic_System_Zoo Probabilistic_Timed_Automata Probabilistic_While Program-Conflict-Analysis Projective_Geometry Projective_Measurements Promela Proof_Strategy_Language PropResPI Propositional_Proof_Systems Prpu_Maxflow PseudoHoops Psi_Calculi Ptolemys_Theorem QHLProver QR_Decomposition Quantales Quaternions Quick_Sort_Cost RIPEMD-160-SPARK ROBDD RSAPSS Ramsey-Infinite Random_BSTs Random_Graph_Subgraph_Threshold Randomised_BSTs Randomised_Social_Choice Rank_Nullity_Theorem Real_Impl Recursion-Addition Recursion-Theory-I Refine_Imperative_HOL Refine_Monadic RefinementReactive Regex_Equivalence Regular-Sets Regular_Algebras Relation_Algebra Relational-Incorrectness-Logic Relational_Disjoint_Set_Forests Relational_Method Relational_Minimum_Spanning_Trees Relational_Paths Rep_Fin_Groups Residuated_Lattices Resolution_FOL Rewriting_Z Ribbon_Proofs Robbins-Conjecture Robinson_Arithmetic Root_Balanced_Tree Routing Roy_Floyd_Warshall SATSolverVerification SC_DOM_Components SDS_Impossibility SIFPL SIFUM_Type_Systems SPARCv8 Safe_Distance Safe_OCL Saturation_Framework Saturation_Framework_Extensions Shadow_DOM Secondary_Sylow Security_Protocol_Refinement Selection_Heap_Sort SenSocialChoice Separata Separation_Algebra Separation_Logic_Imperative_HOL SequentInvertibility Shadow_SC_DOM Shivers-CFA ShortestPath Show Sigma_Commit_Crypto Signature_Groebner Simpl Simple_Firewall Simplex Skew_Heap Skip_Lists Slicing Sliding_Window_Algorithm Smith_Normal_Form Smooth_Manifolds Sort_Encodings Source_Coding_Theorem Special_Function_Bounds Splay_Tree Sqrt_Babylonian Stable_Matching Statecharts Stateful_Protocol_Composition_and_Typing Stellar_Quorums Stern_Brocot Stewart_Apollonius Stirling_Formula Stochastic_Matrices Stone_Algebras Stone_Kleene_Relation_Algebras Stone_Relation_Algebras Store_Buffer_Reduction Stream-Fusion Stream_Fusion_Code Strong_Security Sturm_Sequences Sturm_Tarski Stuttering_Equivalence Subresultants Subset_Boolean_Algebras SumSquares Sunflowers SuperCalc Surprise_Paradox Symmetric_Polynomials Syntax_Independent_Logic Szpilrajn TESL_Language TLA Tail_Recursive_Functions Tarskis_Geometry Taylor_Models Timed_Automata Topological_Semantics Topology TortoiseHare Transcendence_Series_Hancl_Rucki Transformer_Semantics Transition_Systems_and_Automata Transitive-Closure Transitive-Closure-II Treaps Tree-Automata Tree_Decomposition Triangle Trie Twelvefold_Way Tycon Types_Tableaus_and_Goedels_God UPF UPF_Firewall UTP Universal_Turing_Machine UpDown_Scheme Valuation VectorSpace VeriComp Verified-Prover Verified_SAT_Based_AI_Planning VerifyThis2018 VerifyThis2019 Vickrey_Clarke_Groves VolpanoSmith WHATandWHERE_Security WOOT_Strong_Eventual_Consistency WebAssembly Weight_Balanced_Trees Well_Quasi_Orders Winding_Number_Eval Word_Lib WorkerWrapper XML ZFC_in_HOL Zeta_3_Irrational Zeta_Function pGCL