1(* ========================================================================= *)
2(* Create "leakageTheory" setting up the theory of information leakage       *)
3(* ========================================================================= *)
4
5(* ------------------------------------------------------------------------- *)
6(* Load and open relevant theories                                           *)
7(* (Comment out "load" and "loadPath"s for compilation)                      *)
8(* ------------------------------------------------------------------------- *)
9
10(*app load ["bossLib", "metisLib", "arithmeticTheory", "pred_setTheory",
11          "pred_setLib", "listTheory",
12          "state_transformerTheory", "probabilityTheory", "formalizeUseful",
13          "combinTheory", "pairTheory", "realTheory", "realLib", "extra_boolTheory",
14          "jrhUtils", "extra_pred_setTheory", "realSimps", "extra_realTheory",
15          "measureTheory", "numTheory", "simpLib", "seqTheory", "subtypeTheory",
16          "transcTheory", "limTheory", "stringTheory", "rich_listTheory",
17          "stringSimps", "listSimps", "lebesgueTheory",
18          "informationTheory", "extra_stringTheory", "leakageTheory",
19          "extra_stringLib", "leakageLib", "extra_listTheory"];*)
20
21open HolKernel Parse boolLib bossLib;
22
23open metisLib arithmeticTheory pred_setTheory pred_setLib
24     state_transformerTheory
25     formalizeUseful extra_numTheory combinTheory
26     pairTheory realTheory realLib extra_boolTheory jrhUtils
27     extra_pred_setTheory realSimps extra_realTheory numTheory
28     simpLib seqTheory subtypeTheory
29     transcTheory limTheory stringTheory stringSimps
30     informationTheory extra_stringTheory leakageTheory
31     extra_stringLib;
32
33open listTheory rich_listTheory listSimps extra_listTheory;
34open real_sigmaTheory;
35open leakageLib;
36open measureTheory lebesgueTheory probabilityTheory;
37
38(* ------------------------------------------------------------------------- *)
39(* Start a new theory called "information"                                   *)
40(* ------------------------------------------------------------------------- *)
41
42val _ = new_theory "dining_cryptos";
43
44(* ------------------------------------------------------------------------- *)
45(* Helpful proof tools                                                       *)
46(* ------------------------------------------------------------------------- *)
47
48val REVERSE = Tactical.REVERSE;
49val Simplify = RW_TAC arith_ss;
50val Suff = PARSE_TAC SUFF_TAC;
51val Know = PARSE_TAC KNOW_TAC;
52val Rewr = DISCH_THEN (REWRITE_TAC o wrap);
53val Rewr' = DISCH_THEN (ONCE_REWRITE_TAC o wrap);
54
55val Cond = DISCH_THEN (fn mp_th => let
56                           val cond = fst (dest_imp (concl mp_th))
57                       in KNOW_TAC cond >| [ALL_TAC, DISCH_THEN (MP_TAC o MP mp_th)]
58                       end);
59
60val POP_ORW = POP_ASSUM (fn thm => ONCE_REWRITE_TAC [thm]);
61
62val safe_set_ss = bool_ss ++ PRED_SET_ss;
63val set_ss = arith_ss ++ PRED_SET_ss;
64
65(* ************************************************************************* *)
66(* Case Study: The Dining Cryptographers - Definitions                       *)
67(* ************************************************************************* *)
68
69val set_announcements_def = Define
70   `(set_announcements (high: bool state) (low:bool state)
71                       (random:bool state) (n:num) (0:num) (s:string) =
72                if (s = (STRCAT "announces" (toString 0))) then
73                        ((high (STRCAT "pays" (toString 0)))) xor
74                        ((random (STRCAT "coin" (toString 0))) xor
75                         (random (STRCAT "coin" (toString n))))
76                else
77                        low s) /\
78    (set_announcements high low random n (SUC i) s =
79                if (s = (STRCAT "announces" (toString (SUC i)))) then
80                        ((high (STRCAT "pays" (toString (SUC i))))) xor
81                        ((random (STRCAT "coin" (toString (SUC i)))) xor
82                         (random (STRCAT "coin" (toString i))))
83                else
84                        (set_announcements high low random n i) s)`;
85
86val XOR_announces_def = Define
87  `(XOR_announces (low:bool state) (0:num) = low (STRCAT "announces" (toString 0))) /\
88   (XOR_announces low (SUC i) = (low (STRCAT "announces" (toString (SUC i)))) xor
89                                (XOR_announces low i))`;
90
91val compute_result_def = Define
92   `compute_result (low:bool state) (n:num) (s:string) =
93        if (s = "result") then XOR_announces low n else low s`;
94
95val dcprog_def = Define
96   `dcprog (SUC(SUC(SUC n))) = (\((high:bool state, low:bool state), random:bool state).
97        compute_result (set_announcements high low random (SUC(SUC n)) (SUC(SUC n)))
98                       (SUC(SUC n)))`;
99
100val dc_high_states_set_def = Define
101   `(dc_high_states_set (0:num) = {(\s:string. s = (STRCAT "pays" (toString 0)))}) /\
102    (dc_high_states_set (SUC n) = (\s:string. s = (STRCAT "pays" (toString (SUC n))))
103                                                        INSERT (dc_high_states_set n))`;
104
105val dc_high_states_def = Define
106   `dc_high_states nsapays (SUC(SUC n)) = if nsapays then
107                                                {(\s: string. s = STRCAT "pays" (toString (SUC(SUC n))))}
108                                          else
109                                                dc_high_states_set (SUC n)`;
110
111val dc_low_states_def = Define
112   `dc_low_states = {(\s:string. F)}`;
113
114val dc_random_states_def = Define
115  `(dc_random_states (0:num) = {(\s:string. s = (STRCAT "coin" (toString 0))); (\s:string. F)}) /\
116   (dc_random_states (SUC n) = (IMAGE (\s:bool state.
117                                        (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then
118                                                        T
119                                                    else (s x)))
120                                       (dc_random_states n))
121                                    UNION
122                               (IMAGE (\s:bool state.
123                                        (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then
124                                                        F
125                                                    else (s x)))
126                                       (dc_random_states n)))`;
127
128val dc_prog_space_def = Define
129   `dc_prog_space (SUC(SUC n)) nsapays =
130        unif_prog_space
131                (dc_high_states nsapays (SUC(SUC n)))
132                dc_low_states
133                (dc_random_states (SUC n))`;
134
135(* ************************************************************************* *)
136(* Case Study: The Dining Cryptographers - Basic Lemmas                      *)
137(* ************************************************************************* *)
138
139val set_announcements_alt = store_thm  ("set_announcements_alt",
140   ``!high low random n i.
141        (set_announcements high low random n 0 =
142         (\s. if (s = (STRCAT "announces" (toString 0))) then
143                        ((high (STRCAT "pays" (toString 0)))) xor
144                        ((random (STRCAT "coin" (toString 0))) xor (random (STRCAT "coin" (toString n))))
145                else
146                        low s)) /\
147        (set_announcements high low random n (SUC i) =
148         (\s. if (s = (STRCAT "announces" (toString (SUC i)))) then
149                        ((high (STRCAT "pays" (toString (SUC i))))) xor
150                        ((random (STRCAT "coin" (toString (SUC i)))) xor (random (STRCAT "coin" (toString i))))
151                else
152                        (set_announcements high low random n i) s))``,
153   RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC std_ss [set_announcements_def]);
154
155val compute_result_alt = store_thm  ("compute_result_alt",
156   ``!low n. compute_result low n = (\s. if (s = "result") then XOR_announces low n else low s)``,
157   RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC std_ss [compute_result_def]);
158
159val dc_high_states_set_finite = store_thm  ("dc_high_states_set_finite",
160   ``!n. FINITE (dc_high_states_set n)``,
161   Induct
162   >> FULL_SIMP_TAC safe_set_ss [dc_high_states_set_def]);
163
164val dc_low_states_finite = prove  (``FINITE dc_low_states``,
165   RW_TAC set_ss [dc_low_states_def]);
166
167val dc_random_states_finite = prove  (``!n. FINITE (dc_random_states n)``,
168   Induct
169   >> FULL_SIMP_TAC safe_set_ss [dc_random_states_def]);
170
171val IN_dc_high_states_set = store_thm  ("IN_dc_high_states_set",
172   ``!n x. x IN (dc_high_states_set n) =
173           (?i. i <= n /\ (x = (\s. s = STRCAT "pays" (toString i))))``,
174   Induct
175   >> RW_TAC set_ss [dc_high_states_set_def]
176   >> METIS_TAC [DECIDE ``!(x:num) y. x <= SUC y = x <= y \/ (x = SUC y)``, STRCAT_toString_inj]);
177
178val IN_dc_random_states = store_thm
179  ("IN_dc_random_states",
180   ``!n s. s IN (dc_random_states n) = !x. (~(?i. i <= n /\
181                                        (x = STRCAT "coin" (toString i)))) ==>
182                                        (~(s x))``,
183   Induct
184   >- (RW_TAC set_ss [dc_random_states_def] >> METIS_TAC [])
185   >> RW_TAC set_ss [Once dc_random_states_def]
186   >> EQ_TAC
187   >- METIS_TAC [DECIDE ``!(i:num) n. i <= n ==> i <= SUC n``, LESS_EQ_REFL]
188   >> RW_TAC std_ss []
189   >> Cases_on `s (STRCAT "coin" (toString (SUC n)))`
190   >- (DISJ1_TAC >> Q.EXISTS_TAC `(\x. if x = STRCAT "coin" (toString (SUC n)) then F else s x)`
191       >> RW_TAC std_ss [FUN_EQ_THM] >- METIS_TAC []
192       >> METIS_TAC [DECIDE ``!(i:num) n. i <= SUC n = ((i = SUC n) \/ i <= n)``])
193   >> DISJ2_TAC >> Q.EXISTS_TAC `s`
194   >> RW_TAC std_ss [FUN_EQ_THM] >- METIS_TAC []
195   >> METIS_TAC [DECIDE ``!(i:num) n. i <= SUC n = ((i = SUC n) \/ i <= n)``]);
196
197val dc_high_states_set_not_empty = store_thm
198  ("dc_high_states_set_not_empty",
199   ``!n. ~((dc_high_states_set n)={})``,
200   Induct >> RW_TAC set_ss [dc_high_states_set_def]);
201
202val dc_low_states_not_empty = store_thm
203  ("dc_low_states_not_empty",
204   ``~(dc_low_states={})``,
205   RW_TAC set_ss [dc_low_states_def]);
206
207val coin_out_of_range_eq_zero_dc_random_states = store_thm
208  ("coin_out_of_range_eq_zero_dc_random_states",
209   ``!n s. s IN (dc_random_states n) ==>
210                !i. (n < i) ==> ~ (s (STRCAT "coin" (toString i)))``,
211   Induct
212   >- (RW_TAC set_ss [dc_random_states_def]
213       >> RW_TAC std_ss [STRCAT_toString_inj]
214       >> POP_ASSUM MP_TAC >> DECIDE_TAC)
215   >> RW_TAC set_ss [dc_random_states_def]
216   >> RW_TAC std_ss [STRCAT_toString_inj, DECIDE ``!i:num n. SUC n < i ==> ~(i = SUC n)``,                                        DECIDE ``!i:num n. SUC n < i ==> n < i``]);
217
218val dc_random_states_not_empty = store_thm
219  ("dc_random_states_not_empty",
220   ``!n. ~((dc_random_states n)={})``,
221   Induct >> RW_TAC set_ss [dc_random_states_def]);
222
223val CROSS_NON_EMPTY_IMP = prove
224   (``!P Q. FINITE P /\ FINITE Q /\ ~(P={}) /\ ~(Q={}) ==> ~(P CROSS Q = {})``,
225    REPEAT STRIP_TAC
226    THEN (MP_TAC o Q.SPEC `P`) SET_CASES
227    THEN RW_TAC std_ss []
228    THEN SPOSE_NOT_THEN STRIP_ASSUME_TAC
229    THEN (MP_TAC o Q.ISPEC `Q:'b->bool`) SET_CASES
230    THEN RW_TAC std_ss []
231    THEN SPOSE_NOT_THEN STRIP_ASSUME_TAC
232    THEN FULL_SIMP_TAC std_ss [Once EXTENSION, NOT_IN_EMPTY, IN_INSERT, IN_CROSS]
233    THEN METIS_TAC [PAIR, FST, SND, PAIR_EQ]);
234
235val CROSS_HLR_NON_EMPTY_IMP = prove
236   (``!h l r. FINITE h /\ FINITE l /\ FINITE r /\ ~(h={}) /\ ~(l={}) /\ ~(r={}) ==> ~((h CROSS l) CROSS r = {})``,
237    METIS_TAC [CROSS_NON_EMPTY_IMP, FINITE_CROSS]);
238
239val dc_states_cross_not_empty = store_thm
240  ("dc_states_cross_not_empty",
241   ``!n. ~((((dc_high_states_set (SUC(SUC n))) CROSS dc_low_states) CROSS (dc_random_states (SUC(SUC n))))={})``,
242   RW_TAC std_ss [CROSS_HLR_NON_EMPTY_IMP, dc_high_states_set_finite, dc_random_states_finite,
243                  dc_low_states_finite, dc_high_states_set_not_empty, dc_random_states_not_empty,
244                  dc_low_states_not_empty]);
245
246val dc_prog_space_F_set_thm = store_thm
247  ("dc_prog_space_F_set_thm",
248   ``!n. dc_prog_space (SUC(SUC n)) F =
249         unif_prog_space (dc_high_states_set (SUC n))
250                         dc_low_states
251                         (dc_random_states (SUC n))``,
252   RW_TAC std_ss [dc_prog_space_def, dc_high_states_def]);
253
254val dc_prog_space_T_set_thm = store_thm
255  ("dc_prog_space_T_set_thm",
256   ``!n. dc_prog_space (SUC(SUC n)) T =
257         unif_prog_space {(\s:string. s = STRCAT "pays" (toString (SUC(SUC n))))}
258                         {(\s:string. F)}
259                         (dc_random_states (SUC n))``,
260   RW_TAC std_ss [dc_prog_space_def, dc_high_states_def, dc_low_states_def]);
261
262(* ************************************************************************* *)
263(* Case Study: The Dining Cryptographers - Computation for 3 cryptos         *)
264(* ************************************************************************* *)
265
266val fun_eq_lem = METIS_PROVE [] ``!y z. (!x. (x = y) = (x = z)) = (y = z)``;
267
268val card_dc_high_states_set3 = store_thm
269  ("card_dc_high_states_set3",
270   ``& (CARD (dc_high_states_set (SUC (SUC 0)))) = 3``,
271   RW_TAC set_ss [dc_high_states_set_def, FUN_EQ_THM, fun_eq_lem, STRCAT_toString_inj]);
272
273val fun_eq_lem5 = METIS_PROVE []
274        ``!x a b P Q. (x = a) \/ ~ (x = b) /\ P \/ Q = (a = b) /\ ((x = a) \/ P)
275                        \/ ~(x = b) /\ ((x = a) \/ P) \/ Q``;
276
277val CARD_dc_set_cross = store_thm
278  ("CARD_dc_set_cross",
279   ``1 / & (CARD (IMAGE (\s. (FST s,SND s,dcprog (SUC (SUC (SUC 0))) s))
280                ((dc_high_states_set (SUC (SUC 0)) CROSS dc_low_states) CROSS
281                   dc_random_states (SUC (SUC 0))))) = 1 / 24``,
282   RW_TAC set_ss [dcprog_def, dc_low_states_def, dc_high_states_def, dc_high_states_set_def, dc_random_states_def,
283                  CROSS_EQNS, compute_result_alt, XOR_announces_def, set_announcements_def, xor_def]
284   >> CONV_TAC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV set_ss [fun_eq_lem])))
285   >> RW_TAC set_ss [CROSS_EQNS, STRCAT_toString_inj]
286   >> CONV_TAC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV set_ss [fun_eq_lem])))
287   >> REPEAT (SIMP_TAC std_ss [FINITE_EMPTY, FINITE_INSERT, Once CARD_DEF]
288              >> CONV_TAC (FIND_CONV ``x IN y`` (IN_CONV (SIMP_CONV std_ss [FUN_EQ_THM]
289                                        THENC SIMP_CONV bool_ss [EQ_IMP_THM]
290                                        THENC SIMP_CONV bool_ss [DISJ_IMP_THM]
291                                        THENC EVAL
292                                        THENC SIMP_CONV bool_ss [LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
293                                        THENC EVAL
294                                        THENC SIMP_CONV bool_ss [fun_eq_lem5, GSYM LEFT_AND_OVER_OR]
295                                        THENC EVAL
296                                        THENC SIMP_CONV bool_ss [FORALL_AND_THM])))
297               >> SIMP_TAC std_ss []));
298
299val CARD_dc_high_states_set = store_thm
300  ("CARD_dc_high_states_set",
301   ``!n. CARD (dc_high_states_set n) = SUC n``,
302   Induct >> RW_TAC set_ss [dc_high_states_set_def, dc_high_states_set_finite, CARD_INSERT,
303                            IN_dc_high_states_set, FUN_EQ_THM, fun_eq_lem, STRCAT_toString_inj]);
304
305val CARD_DISJOINT_UNION = store_thm
306  ("CARD_DISJOINT_UNION",
307   ``!P Q. FINITE P /\ FINITE Q /\ DISJOINT P Q ==>
308                (CARD (P UNION Q) = CARD P + CARD Q)``,
309   REPEAT STRIP_TAC >> ASM_SIMP_TAC std_ss [GSYM CARD_UNION]
310   >> Suff `P INTER Q = {}` >> RW_TAC set_ss [] >> ONCE_REWRITE_TAC [EXTENSION]
311   >> FULL_SIMP_TAC set_ss [DISJOINT_DEF, EXTENSION]);
312
313val CARD_dc_random_states = store_thm
314  ("CARD_dc_random_states",
315   ``!n. CARD (dc_random_states n) = 2 ** (SUC n)``,
316   Induct >> RW_TAC set_ss [dc_random_states_def, EXP, FUN_EQ_THM]
317   >> `DISJOINT (IMAGE (\s x. (x = STRCAT "coin" (toString (SUC n))) \/ s x)
318                        (dc_random_states n))
319                (IMAGE (\s x. ~(x = STRCAT "coin" (toString (SUC n))) /\ s x)
320                        (dc_random_states n))`
321        by (RW_TAC std_ss [DISJOINT_DEF]
322   >> ONCE_REWRITE_TAC [EXTENSION]
323   >> RW_TAC std_ss [NOT_IN_EMPTY, IN_INTER, IN_IMAGE] >> METIS_TAC [])
324   >> RW_TAC std_ss [IMAGE_FINITE, dc_random_states_finite, CARD_DISJOINT_UNION]
325   >> `CARD (IMAGE (\s x. (x = STRCAT "coin" (toString (SUC n))) \/ s x)
326                        (dc_random_states n)) = CARD (dc_random_states n)`
327        by (MATCH_MP_TAC CARD_IMAGE
328            >> Q.EXISTS_TAC `(IMAGE (\s x. (x = STRCAT "coin" (toString (SUC n))) \/ s x)
329                                        (dc_random_states n))`
330            >> RW_TAC set_ss [dc_random_states_finite, INJ_DEF, FUN_EQ_THM, fun_eq_lem]
331            >- (Q.EXISTS_TAC `s'` >> RW_TAC std_ss [])
332            >> METIS_TAC [coin_out_of_range_eq_zero_dc_random_states, DECIDE ``!n:num. n < SUC n``])
333   >> `CARD (IMAGE (\s x. ~(x = STRCAT "coin" (toString (SUC n))) /\ s x) (dc_random_states n)) =
334       CARD (dc_random_states n)`
335        by (MATCH_MP_TAC CARD_IMAGE
336            >> Q.EXISTS_TAC `(IMAGE (\s x. ~(x = STRCAT "coin" (toString (SUC n))) /\ s x) (dc_random_states n))`
337            >> RW_TAC set_ss [dc_random_states_finite, INJ_DEF, FUN_EQ_THM, fun_eq_lem]
338            >- (Q.EXISTS_TAC `s'` >> RW_TAC std_ss [])
339            >> METIS_TAC [coin_out_of_range_eq_zero_dc_random_states, DECIDE ``!n:num. n < SUC n``])
340   >> RW_TAC arith_ss [EXP]);
341
342val CARD_dc_low_states = store_thm
343  ("CARD_dc_low_states",
344   ``CARD dc_low_states = 1``,
345   RW_TAC set_ss [dc_low_states_def]);
346
347val dc_states3_cross_not_empty = store_thm
348  ("dc_states3_cross_not_empty",
349   ``~((((dc_high_states_set (SUC(SUC 0))) CROSS dc_low_states) CROSS (dc_random_states (SUC(SUC 0))))={})``,
350   ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [dc_high_states_set_def, dc_low_states_def, dc_random_states_def, CROSS_EQNS]
351   >> Q.EXISTS_TAC `(((\s. s = STRCAT "pays" (toString 0)),(\s.F)),(\s.F))`
352   >> RW_TAC std_ss [FST,SND]);
353
354val fun_eq_lem6 = METIS_PROVE []
355  ``!x a b P Q. (x = a) \/ ~ (x = b) /\ P =
356        (a = b) /\ ((x = a) \/ P) \/ ~(x = b) /\ ((x = a) \/ P)``;
357
358val dc3_leakage_result = store_thm
359  ("dc3_leakage_result",
360   ``leakage (dc_prog_space (SUC (SUC (SUC 0))) F) (dcprog (SUC(SUC(SUC 0)))) = 0``,
361   RW_TAC bool_ss [dc_prog_space_F_set_thm, unif_prog_space_leakage_computation_reduce,
362                   dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
363                   dc_states3_cross_not_empty, CARD_dc_high_states_set, CARD_dc_low_states,
364                   CARD_dc_random_states]
365   >> RW_TAC set_ss [dcprog_def, dc_low_states_def, dc_high_states_def, dc_high_states_set_def,
366                     dc_random_states_def, CROSS_EQNS, compute_result_alt, XOR_announces_def,
367                     set_announcements_def, xor_def, STRCAT_toString_inj]
368   >> CONV_TAC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV set_ss [fun_eq_lem])))
369   >> RW_TAC set_ss [CROSS_EQNS, STRCAT_toString_inj]
370   >> CONV_TAC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV set_ss [fun_eq_lem])))
371   >> CONV_TAC
372        (REPEATC ((SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]) THENC
373                  ((ONCE_FIND_CONV ``x DELETE y``
374                     (DELETE_CONV
375                       ((SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND,
376                                             EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM])
377                        THENC EVAL
378                        THENC (SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR,
379                                                  FORALL_AND_THM, DISJ_IMP_THM])
380                        THENC EVAL
381                        THENC (REPEATC (T_F_UNCHANGED_CONV
382                                         ((SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR,
383                                                              FORALL_AND_THM])
384                                          THENC EVAL)))))))
385                  THENC SIMP_CONV arith_ss []))
386   >> CONV_TAC (ONCE_FIND_CONV ``if (x=y) then (1:real) else 0``
387                 (RATOR_CONV (RATOR_CONV (RAND_CONV
388        (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
389         THENC EVAL
390         THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
391         THENC EVAL
392         THENC (REPEATC (T_F_UNCHANGED_CONV
393         (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
394         THENC EVAL))))))
395                THENC SIMP_CONV bool_ss []))
396   >> RW_TAC real_ss [REAL_ADD_ASSOC, GSYM REAL_NEG_ADD, GSYM REAL_ADD_RDISTRIB, REAL_MUL_ASSOC]);
397
398(* ************************************************************************* *)
399(* Case Study: The Dining Cryptographers - Updated Computation for 3 cryptos *)
400(* ************************************************************************* *)
401
402val hl_dups_lemma = prove
403   (``!(s1:string)(s2:string). ((\x. x = s1) = (\x. x = s2)) = (s1=s2)``,
404    METIS_TAC []);
405
406val fun_eq_lem6 = METIS_PROVE []
407        ``!x a b P Q. (x = a) \/ ~ (x = b) /\ P = (a = b) /\ ((x = a) \/ P) \/ ~(x = b) /\ ((x = a) \/ P)``;
408
409val dc_dup_conv = SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
410                                              THENC EVAL
411                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
412                                              THENC EVAL
413                                              THENC (REPEATC (T_F_UNCHANGED_CONV
414                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
415                                                         THENC EVAL)));
416
417val dc_hl_dup_conv = SIMP_CONV arith_ss [hl_dups_lemma, STRCAT_toString_inj];
418
419fun dc_input_unroll (remove_dups_conv:Abbrev.term->Abbrev.thm) =
420        (REPEATC (SIMP_CONV bool_ss [dc_high_states_set_def, dc_low_states_def, dc_random_states_def,
421                                     CARD_DEF, FINITE_INSERT, FINITE_EMPTY, FINITE_SING,
422                                     IMAGE_UNION, IMAGE_IMAGE, combinTheory.o_DEF, IMAGE_INSERT, IMAGE_EMPTY,
423                                     INSERT_UNION, UNION_EMPTY, IN_UNION]
424         THENC (FIND_CONV ``x IN y`` (IN_CONV remove_dups_conv)
425                THENC SIMP_CONV bool_ss [])
426         THENC SIMP_CONV arith_ss []));
427
428
429val new_dc3_leakage_result = store_thm
430  ("new_dc3_leakage_result",
431   ``leakage (dc_prog_space (SUC (SUC (SUC 0))) F) (dcprog (SUC(SUC(SUC 0)))) = 0``,
432   CONV_TAC (SIMP_CONV bool_ss [dc_prog_space_F_set_thm] THENC
433             RATOR_CONV (RAND_CONV (
434             LEAKAGE_COMPUTE_CONV (``dc_high_states_set (SUC(SUC 0))``,``dc_low_states``,``dc_random_states (SUC(SUC 0))``)
435                                  [dc_high_states_set_def, dc_low_states_def, dc_random_states_def]
436                                  [dcprog_def, compute_result_alt, XOR_announces_def, set_announcements_alt, FST, SND, xor_def]
437                                  (dc_input_unroll dc_hl_dup_conv) (dc_input_unroll dc_hl_dup_conv) (dc_input_unroll dc_dup_conv)
438                                  dc_hl_dup_conv dc_hl_dup_conv dc_dup_conv dc_dup_conv))
439             THENC SIMP_CONV real_ss [REAL_ADD_ASSOC, GSYM REAL_NEG_ADD, GSYM REAL_ADD_RDISTRIB, REAL_MUL_ASSOC]));
440
441(* ************************************************************************* *)
442(* Case Study: The Dining Cryptographers - Proofs                            *)
443(* ************************************************************************* *)
444
445(* ************************************************************************* *)
446(* Case Study: The Dining Cryptographers - Proofs (Correctness)              *)
447(* ************************************************************************* *)
448
449val dc_set_announcements_result1 = store_thm
450  ("dc_set_announcements_result1",
451   ``!h l r n. !i. i <= (SUC (SUC n)) ==>
452                ((set_announcements h l r (SUC (SUC n)) i (STRCAT "announces" (toString 0)) =
453                 ((h (STRCAT "pays" (toString 0)))) xor
454                        ((r (STRCAT "coin" (toString 0))) xor (r (STRCAT "coin" (toString (SUC (SUC n))))))))``,
455   NTAC 4 STRIP_TAC >> Induct
456   >> ASM_SIMP_TAC arith_ss [set_announcements_def, STRCAT_toString_inj]);
457
458val dc_set_announcements_result2 = store_thm
459  ("dc_set_announcements_result2",
460   ``!h l r n. ((set_announcements h l r (SUC (SUC n)) (SUC (SUC n)) (STRCAT "announces" (toString 0)) =
461              ((h (STRCAT "pays" (toString 0)))) xor
462                        ((r (STRCAT "coin" (toString 0))) xor (r (STRCAT "coin" (toString (SUC (SUC n))))))))``,
463   ASM_SIMP_TAC arith_ss [dc_set_announcements_result1]);
464
465val dc_set_announcements_result3 = store_thm
466  ("dc_set_announcements_result3",
467   ``!h l r n m i. (SUC i) <= m /\ m < (SUC (SUC (SUC n))) ==>
468                  ((set_announcements h l r (SUC (SUC n)) m) (STRCAT "announces" (toString (SUC i))) =
469                   ((h (STRCAT "pays" (toString (SUC i))))) xor
470                        ((r (STRCAT "coin" (toString (SUC i)))) xor (r (STRCAT "coin" (toString i)))))``,
471   NTAC 4 STRIP_TAC >> Induct
472   >- RW_TAC arith_ss []
473   >> STRIP_TAC >> Cases_on `i = m`
474   >> ASM_SIMP_TAC arith_ss [set_announcements_def, STRCAT_toString_inj]);
475
476val dc_set_announcements_result4 = store_thm
477  ("dc_set_announcements_result4",
478   ``!h l r n i. (SUC i) < (SUC (SUC (SUC n))) ==>
479                  ((set_announcements h l r (SUC (SUC n)) (SUC (SUC n))) (STRCAT "announces" (toString (SUC i))) =
480                   ((h (STRCAT "pays" (toString (SUC i))))) xor
481                        ((r (STRCAT "coin" (toString (SUC i)))) xor (r (STRCAT "coin" (toString i)))))``,
482   ASM_SIMP_TAC arith_ss [dc_set_announcements_result3]);
483
484val dc_set_announcements_result5 = store_thm
485  ("dc_set_announcements_result5",
486   ``!h l r n m s. m <= SUC (SUC n) /\ (!i. i <= m ==> ~(s = STRCAT "announces" (toString i))) ==>
487                 ((set_announcements h l r (SUC (SUC n)) m) s = l s)``,
488   NTAC 4 STRIP_TAC >> Induct
489   >- (STRIP_TAC >> ASM_SIMP_TAC arith_ss [set_announcements_def, STRCAT_toString_inj])
490   >> STRIP_TAC >> ASM_SIMP_TAC arith_ss [set_announcements_def, STRCAT_toString_inj]);
491
492val dc_set_announcements_result6 = store_thm
493  ("dc_set_announcements_result6",
494   ``!h l r n. ((set_announcements h l r (SUC (SUC n)) (SUC (SUC n))) (STRCAT "announces" (toString (SUC (SUC n)))) =
495                   ((h (STRCAT "pays" (toString (SUC (SUC n)))))) xor
496                        ((r (STRCAT "coin" (toString (SUC (SUC n))))) xor (r (STRCAT "coin" (toString (SUC n))))))``,
497   ASM_SIMP_TAC arith_ss [dc_set_announcements_result4]);
498
499(* ------------------------------------------------------------------------- *)
500
501val dc_XOR_announces_result1 = store_thm
502  ("dc_XOR_announces_result1",
503   ``!high low random n i. i <= SUC (SUC n) /\ (high (STRCAT "pays" (toString i))) /\ (!j. (~(j = i)) ==> ~(high (STRCAT "pays" (toString j)))) ==>
504        !k. k < i ==>
505                (XOR_announces (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) k =
506                 ((random (STRCAT "coin" (toString k))) xor (random (STRCAT "coin" (toString (SUC (SUC n)))))))``,
507   NTAC 6 (STRIP_TAC)
508   >> Induct
509   >- (SIMP_TAC bool_ss [XOR_announces_def, dc_set_announcements_result2]
510       >> POP_ASSUM (ASSUME_TAC o Q.SPEC `0`)
511       >> RW_TAC arith_ss [DECIDE ``!n:num. 0 < n = ~(n = 0)``, xor_F])
512   >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
513   >> RW_TAC arith_ss [XOR_announces_def, dc_set_announcements_result4, xor_F]
514   >> METIS_TAC [DECIDE ``!(n:num) m. n < m ==> ~(n = m)``, xor_def]);
515
516val dc_XOR_announces_result2 = store_thm
517  ("dc_XOR_announces_result2",
518   ``!high low random n i. i <= SUC (SUC n) /\ (high (STRCAT "pays" (toString i))) /\ (!j. (~(j = i)) ==> ~(high (STRCAT "pays" (toString j)))) ==>
519        (XOR_announces (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) i =
520                 ~((random (STRCAT "coin" (toString i))) xor (random (STRCAT "coin" (toString (SUC (SUC n)))))))``,
521   REPEAT STRIP_TAC >> Cases_on `i`
522   >- (SIMP_TAC bool_ss [XOR_announces_def, dc_set_announcements_result2]
523       >> RW_TAC bool_ss [EQ_SYM_EQ, xor_def])
524   >> ONCE_REWRITE_TAC [XOR_announces_def]
525   >> (ASSUME_TAC o REWRITE_RULE [DECIDE ``!n:num. n < SUC n``] o Q.SPEC `n'` o UNDISCH_ALL o REWRITE_RULE [GSYM AND_IMP_INTRO]
526       o Q.SPECL [`high`, `low`, `random`, `n`, `(SUC n')`]) dc_XOR_announces_result1
527   >> FULL_SIMP_TAC bool_ss [DECIDE ``!(n:num) m. n <= m = n < SUC m``, dc_set_announcements_result4, DECIDE ``!(n:num) m. SUC n < m ==> n < m``]
528   >> RW_TAC bool_ss [xor_def] >> DECIDE_TAC);
529
530val dc_XOR_announces_result3 = store_thm
531  ("dc_XOR_announces_result3",
532   ``!high low random n i. i <= SUC (SUC n) /\ (high (STRCAT "pays" (toString i))) /\ (!j. (~(j = i)) ==> ~(high (STRCAT "pays" (toString j)))) ==>
533        !k. i <= k /\ k <= SUC (SUC n) ==>
534                (XOR_announces (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) k =
535                 ~((random (STRCAT "coin" (toString k))) xor (random (STRCAT "coin" (toString (SUC (SUC n)))))))``,
536   NTAC 6 (STRIP_TAC) >> Induct
537   >- (RW_TAC arith_ss [XOR_announces_def, dc_set_announcements_result2]
538       >> RW_TAC bool_ss [T_xor])
539   >> Cases_on `SUC k = i`
540   >- ((ASSUME_TAC o UNDISCH_ALL o REWRITE_RULE [GSYM AND_IMP_INTRO] o Q.SPECL [`high`, `low`, `random`, `n`, `i`]) dc_XOR_announces_result2
541       >> ASM_REWRITE_TAC [])
542   >> Cases_on `k = SUC n`
543   >- (ONCE_REWRITE_TAC [XOR_announces_def]
544       >> `i <= SUC n /\ SUC n <= SUC (SUC n)` by FULL_SIMP_TAC arith_ss [] >> FULL_SIMP_TAC bool_ss []
545       >> ASM_SIMP_TAC arith_ss [dc_set_announcements_result6, xor_F]
546       >> RW_TAC bool_ss [xor_def] >> DECIDE_TAC)
547   >> STRIP_TAC >> ONCE_REWRITE_TAC [XOR_announces_def]
548   >> `i <= k /\ k <= SUC (SUC n)` by FULL_SIMP_TAC arith_ss [] >> FULL_SIMP_TAC bool_ss []
549   >> ASM_SIMP_TAC arith_ss [dc_set_announcements_result4, xor_F]
550   >> RW_TAC bool_ss [xor_def] >> DECIDE_TAC);
551
552val dc_XOR_announces_result4 = store_thm
553  ("dc_XOR_announces_result4",
554   ``!high low random n i. i <= SUC (SUC n) /\ (high (STRCAT "pays" (toString i))) /\ (!j. (~(j = i)) ==> ~(high (STRCAT "pays" (toString j)))) ==>
555                (XOR_announces (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) (SUC (SUC n)))``,
556   NTAC 6 STRIP_TAC
557   >> RW_TAC std_ss [(REWRITE_RULE [LESS_EQ_REFL] o UNDISCH o Q.SPEC `SUC (SUC n)`
558                      o UNDISCH_ALL o REWRITE_RULE [GSYM AND_IMP_INTRO]
559                      o Q.SPECL [`high`, `low`, `random`, `n`, `i`]) dc_XOR_announces_result3, xor_refl]);
560
561val dc_XOR_announces_result5 = store_thm
562  ("dc_XOR_announces_result5",
563   ``!high low random n. high IN dc_high_states F (SUC (SUC (SUC n)))  ==>
564                (XOR_announces (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) (SUC (SUC n)))``,
565   NTAC 5 STRIP_TAC
566   >> MATCH_MP_TAC dc_XOR_announces_result4
567   >> FULL_SIMP_TAC bool_ss [dc_high_states_def, IN_dc_high_states_set, GSPECIFICATION, PAIR_EQ]
568   >> Q.EXISTS_TAC `i`
569   >> RW_TAC bool_ss []
570   >> RW_TAC std_ss [STRCAT_toString_inj]);
571
572(* ------------------------------------------------------------------------- *)
573
574val dcprog_result1 = store_thm
575  ("dcprog_result1",
576   ``!high low random n. dcprog (SUC (SUC (SUC n))) ((high, low), random) "result" =
577                         XOR_announces (set_announcements high low random
578                                        (SUC (SUC n)) (SUC (SUC n))) (SUC (SUC n))``,
579   RW_TAC std_ss [dcprog_def, compute_result_def]);
580
581val dcprog_result2 = store_thm
582  ("dcprog_result2",
583   ``!high low random n x. (~(x = "result")) ==>
584        (dcprog (SUC (SUC (SUC n))) ((high, low), random) x =
585         (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) x)``,
586   RW_TAC std_ss [dcprog_def, compute_result_def]);
587
588val dcprog_result3 = store_thm
589  ("dcprog_result3",
590   ``!high low random n i. XOR_announces (dcprog (SUC (SUC (SUC n))) ((high, low), random)) i =
591                         XOR_announces (set_announcements high low random
592                                        (SUC (SUC n)) (SUC (SUC n))) i``,
593   NTAC 4 STRIP_TAC >> Induct
594   >- RW_TAC string_ss [XOR_announces_def, dcprog_result2]
595   >> ONCE_REWRITE_TAC [XOR_announces_def]
596   >> RW_TAC string_ss [dcprog_result2]);
597
598val dcprog_result4 = store_thm
599  ("dcprog_result4",
600   ``!high low random n. XOR_announces (dcprog (SUC (SUC (SUC n))) ((high, low), random)) (SUC (SUC n)) =
601                         (dcprog (SUC (SUC (SUC n))) ((high, low), random)) "result"``,
602   REPEAT STRIP_TAC >> ONCE_REWRITE_TAC [dcprog_result3]
603   >> RW_TAC std_ss [dcprog_def, compute_result_def]);
604
605val dcprog_result5 = store_thm
606  ("dcprog_result5",
607   ``!high low random n. high IN dc_high_states F (SUC (SUC (SUC n))) ==>
608        ((dcprog (SUC (SUC (SUC n))) ((high, low), random)) "result")``,
609   RW_TAC std_ss [dcprog_def, compute_result_def, dc_XOR_announces_result5]);
610
611val dcprog_result6 = store_thm
612  ("dcprog_result6",
613   ``!high low random n i.
614        (dcprog (SUC (SUC (SUC n))) ((high, low), random) (STRCAT "announces" (toString i)) =
615         (set_announcements high low random (SUC (SUC n)) (SUC (SUC n))) (STRCAT "announces" (toString i)))``,
616   RW_TAC string_ss [dcprog_result2]);
617
618(* ************************************************************************* *)
619(* Case Study: The Dining Cryptographers - Proofs (Anonymity)                *)
620(* ************************************************************************* *)
621
622val dc_valid_outputs_def = Define
623   `dc_valid_outputs n = {s: bool state | (s "result" = XOR_announces s n) /\
624                                          (XOR_announces s n) /\
625                                          (!x. (~(x = "result")) /\
626                                               (!i:num. i <= n ==>
627                                                        ~(x = (STRCAT "announces" (toString i)))) ==>
628                                                ~ s x)}`;
629
630val dc_random_witness_def = Define
631   `(dc_random_witness x (0:num) = (\s:string. if s = (STRCAT "coin" (toString 0)) then
632                                                        ~(x (STRCAT "announces" (toString 0)))
633                                                   else F)) /\
634    (dc_random_witness x (SUC i) = (\s. if s = (STRCAT "coin" (toString (SUC i))) then
635                                                ~ (XOR_announces x (SUC i))
636                                            else
637                                                dc_random_witness x i s))`;
638
639val dc_random_witness_lem0 = prove
640  (``!x n i. (SUC n) <= i ==> ~(dc_random_witness x n (STRCAT "coin" (toString i)))``,
641   STRIP_TAC >> Induct >> RW_TAC arith_ss [dc_random_witness_def, STRCAT_toString_inj]);
642
643val dc_random_witness_lem1 = prove
644  (``!x n. ~(dc_random_witness x n (STRCAT "coin" (toString (SUC n))))``,
645   ASM_SIMP_TAC arith_ss [dc_random_witness_lem0]);
646
647val dc_random_witness_lem2 = prove
648  (``!x n. (dc_random_witness x n (STRCAT "coin" (toString 0))) = ~(x (STRCAT "announces" (toString 0)))``,
649   STRIP_TAC >> Induct >> RW_TAC arith_ss [dc_random_witness_def, STRCAT_toString_inj]);
650
651val dc_random_witness_lem3 = prove
652  (``!x n i. i <= n ==> ((dc_random_witness x n (STRCAT "coin" (toString i))) =
653                         ~ (XOR_announces x i))``,
654   STRIP_TAC >> Induct
655   >- RW_TAC std_ss [dc_random_witness_lem2, XOR_announces_def]
656   >> RW_TAC std_ss [dc_random_witness_def, STRCAT_toString_inj]
657   >> FULL_SIMP_TAC arith_ss []);
658
659val dc_random_witness_lem4 = prove
660  (``!x x' n. (!i. i <= n ==> ~(x' = (STRCAT "coin" (toString i)))) ==>
661                ~ dc_random_witness x n x'``,
662   STRIP_TAC >> STRIP_TAC >> Induct
663   >> RW_TAC std_ss [dc_random_witness_def]
664   >> FULL_SIMP_TAC std_ss [DECIDE ``!(n:num) m. n <= m ==> n <= SUC m``]);
665
666val dc_random_witness_result = prove
667  (``!x n i. i <= (SUC(SUC n)) /\ XOR_announces x (SUC(SUC n)) ==>
668             (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString 0)),(\s. F)),dc_random_witness x (SUC n))
669                (STRCAT "announces" (toString i)) =
670              x (STRCAT "announces" (toString i)))``,
671   REPEAT STRIP_TAC >> Cases_on `i`
672   >- RW_TAC arith_ss [dcprog_result6, dc_set_announcements_result2, dc_random_witness_lem1, T_xor, xor_F, dc_random_witness_lem2, xor_comm]
673   >> RW_TAC arith_ss [dcprog_result6, dc_set_announcements_result4, STRCAT_toString_inj, xor_F, dc_random_witness_lem3]
674   >> (ASSUME_TAC o Q.SPECL [`x`, `SUC n`, `SUC n'`]) dc_random_witness_lem3
675   >> Cases_on `n' = SUC n`
676   >- (FULL_SIMP_TAC arith_ss [dc_random_witness_lem1, xor_F] >> METIS_TAC [XOR_announces_def, xor_def])
677   >> FULL_SIMP_TAC arith_ss [] >> ONCE_REWRITE_TAC [XOR_announces_def]
678   >> RW_TAC bool_ss [xor_def] >> DECIDE_TAC);
679
680val dc_valid_outputs_eq_outputs = store_thm
681  ("dc_valid_outputs_eq_outputs",
682   ``!n. (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
683               (dc_high_states F (SUC (SUC (SUC n))) CROSS
684                dc_random_states (SUC (SUC n)))) =
685          dc_valid_outputs (SUC (SUC n))``,
686   RW_TAC std_ss [EXTENSION, IN_IMAGE, IN_CROSS, dc_valid_outputs_def, GSPECIFICATION]
687   >> RW_TAC std_ss [GSYM EXTENSION]
688   >> EQ_TAC
689   >- (STRIP_TAC >> (ASSUME_TAC o Q.ISPEC `(x' :bool state # bool state)`) pair_CASES
690       >> FULL_SIMP_TAC std_ss [dcprog_result4]
691       >> FULL_SIMP_TAC std_ss [FST, SND, dcprog_result5]
692       >> RW_TAC std_ss [dcprog_def, compute_result_def]
693       >> ASM_SIMP_TAC arith_ss [dc_set_announcements_result5])
694   >> STRIP_TAC
695   >> Q.EXISTS_TAC `((\s:string. s = (STRCAT "pays" (toString 0))), dc_random_witness x (SUC n))`
696   >> `(\s. s = STRCAT "pays" (toString 0)) IN dc_high_states F (SUC (SUC (SUC n)))`
697        by (RW_TAC std_ss [dc_high_states_def, IN_dc_high_states_set] >> Q.EXISTS_TAC `0` >> RW_TAC std_ss [])
698   >> `(dc_random_witness x (SUC n)) IN (dc_random_states (SUC (SUC n)))`
699        by (RW_TAC std_ss [IN_dc_random_states]
700            >> POP_ASSUM MP_TAC
701            >> SPEC_TAC (``n:num``,``n:num``)
702            >> Induct
703            >- (RW_TAC bool_ss [dc_random_witness_def]
704                >- (POP_ASSUM (ASSUME_TAC o Q.SPEC `SUC 0`) >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj])
705                >> POP_ASSUM (ASSUME_TAC o Q.SPEC `0`) >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj])
706            >> STRIP_TAC
707            >> `(!i. ~(i <= SUC (SUC n')) \/ ~(x' = STRCAT "coin" (toString i)))`
708                by (STRIP_TAC >> REVERSE (Cases_on `i <= SUC (SUC n')`) >- ASM_REWRITE_TAC []
709                    >> Q.PAT_X_ASSUM `!i. ~(i <= SUC (SUC (SUC n'))) \/ ~(x' = STRCAT "coin" (toString i))`
710                        (ASSUME_TAC o Q.SPEC `i`)
711                    >> FULL_SIMP_TAC arith_ss [])
712            >> FULL_SIMP_TAC bool_ss []
713            >> ONCE_REWRITE_TAC [dc_random_witness_def]
714            >> RW_TAC bool_ss []
715            >> METIS_TAC [LESS_EQ_REFL, STRCAT_toString_inj,
716                          DECIDE ``!(i:num) n. i <= SUC n = ((i = SUC n) \/ i <= n)``])
717   >> ASM_REWRITE_TAC [FST, SND]
718   >> RW_TAC std_ss [FUN_EQ_THM]
719   >> Q.PAT_X_ASSUM `!x'. b` (ASSUME_TAC o Q.SPEC `x'`)
720   >> Cases_on `x' = "result"`
721   >- (ASM_REWRITE_TAC [] >> MATCH_MP_TAC dcprog_result5 >> ASM_REWRITE_TAC [])
722   >> Cases_on `(!i. i <= SUC (SUC n) ==> ~(x' = STRCAT "announces" (toString i)))`
723   >- FULL_SIMP_TAC arith_ss [dcprog_result2, dc_set_announcements_result5]
724   >> FULL_SIMP_TAC bool_ss [dc_random_witness_result]);
725
726(* ------------------------------------------------------------------------- *)
727
728val n_minus_1_announces_list = Define
729   `(n_minus_1_announces_list (0:num) = [(\s:string. s = (STRCAT "announces" (toString 0))); (\s:string. F)]) /\
730    (n_minus_1_announces_list (SUC n) = (MAP (\s:bool state.
731                                              (\x:string. if x = (STRCAT "announces" (toString (SUC n))) then
732                                                                T
733                                                          else s x))
734                                          (n_minus_1_announces_list n)) >>
735                                     (MAP (\s:bool state.
736                                              (\x:string. if x = (STRCAT "announces" (toString (SUC n))) then
737                                                                F
738                                                          else s x))
739                                          (n_minus_1_announces_list n)))`;
740
741val dc_valid_outputs_list = Define
742   `dc_valid_outputs_list (SUC(SUC(SUC n))) =
743        MAP (\l s. (s = "result") \/
744                   (if (s = (STRCAT "announces" (toString (SUC(SUC n))))) then
745                        ~(XOR_announces l (SUC n))
746                    else
747                        l s))
748             (n_minus_1_announces_list (SUC n))`;
749
750val MEM_n_minus_1_announces_list = prove
751  (``!n x. MEM x (n_minus_1_announces_list n) =
752           (!s. (!i. i <= n ==> ~(s = STRCAT "announces" (toString i))) ==> ~x s)``,
753   Induct
754   >- (RW_TAC std_ss [n_minus_1_announces_list, MEM] >> METIS_TAC [])
755   >> RW_TAC list_ss [n_minus_1_announces_list, MEM_MAP]
756   >> EQ_TAC
757   >- (RW_TAC std_ss [] >> RW_TAC std_ss [] >> Q.PAT_X_ASSUM `!s'. b ==> ~s s'` MATCH_MP_TAC >> STRIP_TAC >> FULL_SIMP_TAC arith_ss [])
758   >> RW_TAC std_ss [] >> Cases_on `x (STRCAT "announces" (toString (SUC n)))`
759   >- (DISJ1_TAC >> Q.EXISTS_TAC `(\x'. if (x' = STRCAT "announces" (toString (SUC n))) then F else x x')`
760       >> RW_TAC std_ss [] >- (FULL_SIMP_TAC string_ss [] >> METIS_TAC [])
761       >> Cases_on `(s' = STRCAT "announces" (toString (SUC n)))` >> ASM_REWRITE_TAC []
762       >> Q.PAT_X_ASSUM `!s. b ==> ~x s` MATCH_MP_TAC
763       >> RW_TAC std_ss [DECIDE ``!(n:num) m. n <= SUC m = ((n = SUC m) \/ n < SUC m)``] >> FULL_SIMP_TAC string_ss [])
764   >> DISJ2_TAC >> Q.EXISTS_TAC `x` >> RW_TAC std_ss [] >- (FULL_SIMP_TAC string_ss [] >> METIS_TAC [])
765   >> Cases_on `(s' = STRCAT "announces" (toString (SUC n)))` >> ASM_REWRITE_TAC []
766   >> Q.PAT_X_ASSUM `!s. b ==> ~x s` MATCH_MP_TAC
767   >> RW_TAC std_ss [DECIDE ``!(n:num) m. n <= SUC m = ((n = SUC m) \/ n < SUC m)``] >> FULL_SIMP_TAC string_ss []);
768
769val dc_valid_outputs_eq_dc_valid_outputs_list = prove
770  (``!n. dc_valid_outputs (SUC(SUC n)) = set (dc_valid_outputs_list (SUC(SUC(SUC n))))``,
771   REPEAT STRIP_TAC
772   >> SIMP_TAC std_ss [EXTENSION, IN_LIST_TO_SET, dc_valid_outputs_def, GSPECIFICATION]
773   >> SIMP_TAC std_ss [GSYM EXTENSION, dc_valid_outputs_list, MEM_MAP, MEM_n_minus_1_announces_list]
774   >> STRIP_TAC
775   >> EQ_TAC
776   >- (RW_TAC std_ss [] >> Q.EXISTS_TAC `(\s. if s = "result" then F else
777                                                if s = STRCAT "announces" (toString (SUC (SUC n))) then F
778                                                else  x s)` >> CONJ_TAC
779       >- (`x (STRCAT "announces" (toString (SUC (SUC n)))) = ~XOR_announces (\s. if s = "result" then F else
780                                                if s = STRCAT "announces" (toString (SUC (SUC n))) then F
781                                                else  x s) (SUC n)`
782                by (Q.PAT_X_ASSUM `XOR_announces x (SUC (SUC n))` MP_TAC
783                    >> CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [XOR_announces_def]))
784                    >> Suff `!k n. n <= k ==>
785                                 (XOR_announces x (SUC n) = XOR_announces (\s. if s = "result" then F else
786                                                if s = STRCAT "announces" (toString (SUC (SUC k))) then F
787                                                else  x s) (SUC n))`
788                    >- METIS_TAC [xor_def, LESS_EQ_REFL]
789                    >> STRIP_TAC >> Induct
790                    >- RW_TAC string_ss [XOR_announces_def, STRCAT_toString_inj]
791                    >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
792                    >> ONCE_REWRITE_TAC [XOR_announces_def]
793                    >> ASM_REWRITE_TAC []
794                    >> RW_TAC string_ss [STRCAT_toString_inj])
795           >> METIS_TAC [])
796       >> STRIP_TAC >> STRIP_TAC >> SIMP_TAC std_ss [] >> Cases_on `s = "result"` >> ASM_REWRITE_TAC []
797       >> Cases_on `s = STRCAT "announces" (toString (SUC (SUC n)))` >> ASM_REWRITE_TAC []
798       >> Q.PAT_X_ASSUM `!x'. b ==> ~x x'` MATCH_MP_TAC >> RW_TAC std_ss []
799       >> Cases_on `i = SUC(SUC n)` >> FULL_SIMP_TAC arith_ss [])
800   >> STRIP_TAC
801   >> `XOR_announces x (SUC (SUC n))`
802        by (ASM_REWRITE_TAC [] >> ONCE_REWRITE_TAC [XOR_announces_def] >> SIMP_TAC std_ss []
803            >> `~((STRCAT "announces" (toString (SUC (SUC n))) = "result"))` by SIMP_TAC string_ss [] >> ASM_REWRITE_TAC []
804            >> Suff `!n k. k <= n ==>
805                                (XOR_announces (\s. (s = "result") \/ (if s = STRCAT "announces" (toString (SUC (SUC n))) then
806                                ~XOR_announces l' (SUC n) else l' s)) (SUC k) =
807                                XOR_announces l' (SUC k))`
808            >- RW_TAC arith_ss [xor_inv, xor_comm]
809            >> STRIP_TAC >> Induct >- RW_TAC string_ss [XOR_announces_def, STRCAT_toString_inj]
810            >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
811            >> ONCE_REWRITE_TAC [XOR_announces_def]
812            >> ASM_REWRITE_TAC []
813            >> RW_TAC string_ss [STRCAT_toString_inj])
814   >> ASM_REWRITE_TAC []
815   >> RW_TAC std_ss []
816   >> Q.PAT_X_ASSUM `!s. b ==> ~l' s` MATCH_MP_TAC >> STRIP_TAC >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []);
817
818val LENGTH_n_minus_1_announces_list = prove
819  (``!n. LENGTH (n_minus_1_announces_list n) = 2 ** (SUC n)``,
820   Induct >> RW_TAC list_ss [n_minus_1_announces_list, EXP]);
821
822val LENGTH_dc_valid_outputs_list = prove
823  (``!n. LENGTH (dc_valid_outputs_list (SUC(SUC(SUC n)))) = 2 ** (SUC(SUC n))``,
824   RW_TAC list_ss [dc_valid_outputs_list, LENGTH_n_minus_1_announces_list]);
825
826val n_minus_1_announces_list_SUC_n = prove
827  (``!n k (x:bool state). n <= k /\ MEM x (n_minus_1_announces_list n) ==> ~(x (STRCAT "announces" (toString (SUC k))))``,
828   Induct >- (RW_TAC std_ss [n_minus_1_announces_list]
829              >> FULL_SIMP_TAC list_ss [n_minus_1_announces_list, STRCAT_toString_inj, toString_inj])
830   >> STRIP_TAC >> STRIP_TAC >> SIMP_TAC list_ss [n_minus_1_announces_list, MEM_MAP]
831   >> STRIP_TAC >> ASM_REWRITE_TAC [] >> SIMP_TAC std_ss [STRCAT_toString_inj] >> FULL_SIMP_TAC string_ss [toString_inj]);
832
833val ALL_DISTINCT_n_minus_1_announces_list = prove
834  (``!n. ALL_DISTINCT (n_minus_1_announces_list n)``,
835   Induct >- (RW_TAC list_ss [ALL_DISTINCT, n_minus_1_announces_list] >> METIS_TAC [])
836   >> RW_TAC bool_ss [n_minus_1_announces_list]
837   >> MATCH_MP_TAC ALL_DISTINCT_APPEND
838   >> CONJ_TAC
839   >- (MATCH_MP_TAC ALL_DISTINCT_MAP2 >> RW_TAC bool_ss [FUN_EQ_THM] >> METIS_TAC [n_minus_1_announces_list_SUC_n, LESS_EQ_REFL])
840   >> CONJ_TAC
841   >- (MATCH_MP_TAC ALL_DISTINCT_MAP2 >> RW_TAC bool_ss [FUN_EQ_THM] >> METIS_TAC [n_minus_1_announces_list_SUC_n, LESS_EQ_REFL])
842   >> RW_TAC bool_ss [MEM_MAP]
843   >> Cases_on `(!s.
844       ~(x = (\x. (x = STRCAT "announces" (toString (SUC n))) \/ s x)) \/
845       ~MEM s (n_minus_1_announces_list n))`
846   >> ASM_REWRITE_TAC []
847   >> FULL_SIMP_TAC bool_ss []
848   >> RW_TAC bool_ss [] >> REVERSE (Cases_on `MEM s' (n_minus_1_announces_list n)`) >> ASM_REWRITE_TAC []
849   >> RW_TAC bool_ss [FUN_EQ_THM] >> Q.EXISTS_TAC `STRCAT "announces" (toString (SUC n))`
850   >> RW_TAC bool_ss [n_minus_1_announces_list_SUC_n]);
851
852val n_minus_1_announces_list_result = prove
853  (``!n x. MEM x (n_minus_1_announces_list n) ==> ~(x "result")``,
854   Induct >- (RW_TAC list_ss [n_minus_1_announces_list] >> RW_TAC arith_string_ss [])
855   >> SIMP_TAC list_ss [n_minus_1_announces_list, MEM_MAP]
856   >> STRIP_TAC >> STRIP_TAC >> ASM_REWRITE_TAC [] >> SIMP_TAC arith_string_ss [] >> RW_TAC std_ss []);
857
858val ALL_DISTINCT_dc_valid_outputs_list = prove
859  (``!n. ALL_DISTINCT (dc_valid_outputs_list (SUC(SUC(SUC n))))``,
860   STRIP_TAC >> ONCE_REWRITE_TAC [dc_valid_outputs_list] >> MATCH_MP_TAC ALL_DISTINCT_MAP2
861   >> RW_TAC std_ss [ALL_DISTINCT_n_minus_1_announces_list, FUN_EQ_THM]
862   >> POP_ASSUM (ASSUME_TAC o Q.SPEC `x'`)
863   >> METIS_TAC [n_minus_1_announces_list_result, n_minus_1_announces_list_SUC_n, LESS_EQ_REFL]);
864
865val CARD_dc_valid_outputs = store_thm
866  ("CARD_dc_valid_outputs",
867   ``!n. CARD (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
868               (dc_high_states F (SUC (SUC (SUC n))) CROSS
869                dc_random_states (SUC (SUC n)))) =
870          2 ** (SUC(SUC n))``,
871   RW_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_eq_dc_valid_outputs_list, GSYM LENGTH_dc_valid_outputs_list,
872                  CARD_LIST_TO_SET, ALL_DISTINCT_dc_valid_outputs_list, MAKE_ALL_DISTINCT_ALL_DISTINCT]);
873
874(* ------------------------------------------------------------------------- *)
875
876val valid_coin_assignment = Define
877  `(valid_coin_assignment r out h n (0:num) =
878        (r (STRCAT "coin" (toString 0)) =
879                ((r (STRCAT "coin" (toString (SUC (SUC n))))) xor (XOR_announces out (0:num)) xor ~((0:num) < h)))) /\
880    (valid_coin_assignment r out h n (SUC i) =
881        (r (STRCAT "coin" (toString (SUC i))) =
882                ((r (STRCAT "coin" (toString (SUC (SUC n))))) xor (XOR_announces out (SUC i)) xor ~((SUC i) < h))) /\
883        valid_coin_assignment r out h n i)`;
884
885val valid_coin_assignment_result1 = prove
886  (``!x out p n i j. j <= i /\ valid_coin_assignment x out p n i ==> valid_coin_assignment x out p n j``,
887   NTAC 4 STRIP_TAC >> Induct >- RW_TAC arith_ss []
888   >> REPEAT STRIP_TAC >> Cases_on `j = SUC i` >> FULL_SIMP_TAC arith_ss []
889   >> Q.PAT_X_ASSUM `!j. b` MATCH_MP_TAC >> FULL_SIMP_TAC arith_ss [valid_coin_assignment]);
890
891val valid_coin_set_eq_valid_coin_assignment = store_thm
892  ("valid_coin_set_eq_valid_coin_assignment",
893   ``!n p out. p <= SUC (SUC n) /\
894               out IN (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
895                             (dc_high_states F (SUC (SUC (SUC n))) CROSS dc_random_states (SUC (SUC n)))) ==>
896             ({r | r IN (dc_random_states (SUC (SUC n))) /\
897                   (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString p)), (\s.F)), r) = out)} =
898              {r | r IN (dc_random_states (SUC (SUC n))) /\
899                   valid_coin_assignment r out p n (SUC (SUC n))})``,
900   RW_TAC std_ss [EXTENSION, GSPECIFICATION]
901   >> RW_TAC std_ss [GSYM EXTENSION]
902   >> REVERSE (Cases_on `x IN dc_random_states (SUC (SUC n))`) >> ASM_REWRITE_TAC []
903   >> `(\s. s = STRCAT "pays" (toString p)) IN dc_high_states F (SUC (SUC (SUC n)))`
904        by (RW_TAC std_ss [dc_high_states_def, IN_dc_high_states_set] >> Q.EXISTS_TAC `p` >> RW_TAC std_ss [])
905   >> EQ_TAC
906   >- (STRIP_TAC >> POP_ASSUM (ASSUME_TAC o ONCE_REWRITE_RULE [EQ_SYM_EQ])
907       >> Suff `!i. i <= SUC (SUC n) ==> valid_coin_assignment x out p n i` >- RW_TAC std_ss [LESS_EQ_REFL]
908       >> FULL_SIMP_TAC std_ss [] >> Induct
909       >- (RW_TAC std_ss [valid_coin_assignment, XOR_announces_def, dcprog_result6, dc_set_announcements_result2,
910                          STRCAT_toString_inj]
911           >> Cases_on `p = 0` >> RW_TAC arith_ss [T_xor, xor_F, xor_def] >> DECIDE_TAC)
912       >> STRIP_TAC >> FULL_SIMP_TAC arith_ss [] >> ONCE_REWRITE_TAC [valid_coin_assignment]
913       >> ASM_REWRITE_TAC [] >> ONCE_REWRITE_TAC [XOR_announces_def]
914       >> `SUC i < SUC (SUC (SUC n))` by FULL_SIMP_TAC arith_ss []
915       >> `i < SUC (SUC (SUC n))` by FULL_SIMP_TAC arith_ss []
916       >> RW_TAC bool_ss [dcprog_result6, dc_set_announcements_result4, STRCAT_toString_inj]
917       >> Cases_on `i`
918       >- (RW_TAC std_ss [valid_coin_assignment, XOR_announces_def, dcprog_result6, dc_set_announcements_result2,
919                          STRCAT_toString_inj, xor_F]
920           >> Cases_on `p = 0` >- (RW_TAC arith_ss [xor_def] >> DECIDE_TAC)
921           >> Cases_on `p = 1` >- (RW_TAC arith_ss [xor_def] >> DECIDE_TAC)
922           >> FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC)
923       >> FULL_SIMP_TAC bool_ss [valid_coin_assignment]
924       >> RW_TAC bool_ss [xor_F, xor_def] >> Cases_on `p = SUC (SUC n')` >- (FULL_SIMP_TAC arith_ss [] >> DECIDE_TAC)
925       >> FULL_SIMP_TAC arith_ss [] >> Cases_on `SUC (SUC n') < p` >> FULL_SIMP_TAC arith_ss [] >> DECIDE_TAC)
926   >> STRIP_TAC >> RW_TAC std_ss [FUN_EQ_THM] >> Cases_on `x' = "result"`
927   >- (ASM_SIMP_TAC std_ss [dcprog_result5] >> FULL_SIMP_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION])
928   >> Cases_on `!i. i <= SUC (SUC n) ==> ~(x' = STRCAT "announces" (toString i))`
929   >- (ASM_SIMP_TAC std_ss [dcprog_result2, dc_set_announcements_result5, LESS_EQ_REFL]
930       >> FULL_SIMP_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION])
931   >> FULL_SIMP_TAC std_ss [] >> Induct_on `i`
932   >- (RW_TAC std_ss [dcprog_result6, dc_set_announcements_result2, STRCAT_toString_inj]
933       >> `valid_coin_assignment x out p n 0` by METIS_TAC [valid_coin_assignment_result1, ZERO_LESS_EQ]
934       >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def])
935       >> POP_ORW >> Cases_on `p = 0` >> FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC)
936   >> STRIP_TAC >> STRIP_TAC
937   >> `SUC i < SUC (SUC (SUC n))` by FULL_SIMP_TAC arith_ss []
938   >> `i < SUC (SUC (SUC n))` by FULL_SIMP_TAC arith_ss []
939   >> RW_TAC bool_ss [dcprog_result6, dc_set_announcements_result4, STRCAT_toString_inj]
940   >> `valid_coin_assignment x out p n (SUC i)`
941        by (MATCH_MP_TAC valid_coin_assignment_result1 >> Q.EXISTS_TAC `SUC (SUC n)` >> FULL_SIMP_TAC arith_ss [])
942   >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def])
943   >> POP_ORW
944   >> `valid_coin_assignment x out p n i`
945        by (MATCH_MP_TAC valid_coin_assignment_result1 >> Q.EXISTS_TAC `SUC (SUC n)` >> FULL_SIMP_TAC arith_ss [])
946   >> Cases_on `i`
947   >- (RW_TAC std_ss [XOR_announces_def]
948       >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def]) >> POP_ORW
949       >> Cases_on `p = 1` >- (FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC)
950       >> FULL_SIMP_TAC arith_ss [xor_def] >> Cases_on `p = 0` >> FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC)
951   >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def]) >> POP_ORW
952   >> ONCE_REWRITE_TAC [XOR_announces_def]
953   >> Cases_on `p = SUC (SUC n')` >- (FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC)
954   >> FULL_SIMP_TAC arith_ss [xor_def] >> Cases_on `SUC (SUC n') < p` >> FULL_SIMP_TAC arith_ss [xor_def] >> DECIDE_TAC);
955
956val coin_assignment = Define
957   `(coin_assignment out h n choice (0:num) =
958        (\s. if s = (STRCAT "coin" (toString 0)) then
959                (choice xor (XOR_announces out (0:num)) xor ~((0:num) < h))
960             else F)) /\
961    (coin_assignment out h n choice (SUC i) =
962        (\s. if s = (STRCAT "coin" (toString (SUC i))) then
963                (choice xor (XOR_announces out (SUC i)) xor ~((SUC i) < h))
964             else
965                coin_assignment out h n choice i s))`;
966
967val coin_assignment_set = Define
968  `!out p n. coin_assignment_set out p n =
969        {coin_assignment out p n T (SUC(SUC n));
970          coin_assignment out p n F (SUC(SUC n))}`;
971
972val coin_assignment_result1 = prove
973  (``!out h choice x' n i. (!j. j <= i ==> ~(x' = STRCAT "coin" (toString j))) ==> ~(coin_assignment out h n choice i x')``,
974   NTAC 5 STRIP_TAC >> Induct
975   >- RW_TAC arith_ss [coin_assignment, STRCAT_toString_inj]
976   >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
977   >> RW_TAC arith_ss [coin_assignment, STRCAT_toString_inj]);
978
979val coin_assignment_result2 = prove
980  (``!out h choice n j i. i <= j ==>
981        (coin_assignment out p n choice j (STRCAT "coin" (toString i)) =
982         coin_assignment out p n choice i (STRCAT "coin" (toString i)))``,
983   NTAC 4 STRIP_TAC >> Induct >- RW_TAC arith_ss []
984   >> RW_TAC std_ss [coin_assignment, STRCAT_toString_inj]
985   >> FULL_SIMP_TAC arith_ss []);
986
987val valid_coin_assignment_eq_2_element_set = store_thm
988  ("valid_coin_assignment_eq_2_element_set",
989   ``!n p out.
990         p <= SUC (SUC n) /\
991         out IN
992         IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
993           (dc_high_states F (SUC (SUC (SUC n))) CROSS
994            dc_random_states (SUC (SUC n))) ==>
995        ({r |
996           r IN dc_random_states (SUC (SUC n)) /\
997           valid_coin_assignment r out p n (SUC (SUC n))} =
998         coin_assignment_set out p n)``,
999   RW_TAC std_ss [EXTENSION, GSPECIFICATION, coin_assignment_set]
1000   >> RW_TAC std_ss [GSYM EXTENSION, IN_INSERT, NOT_IN_EMPTY]
1001   >> EQ_TAC
1002   >- (STRIP_TAC
1003       >> Cases_on `x (STRCAT "coin" (toString (SUC(SUC n))))`
1004       >- (DISJ1_TAC
1005           >> RW_TAC std_ss [FUN_EQ_THM]
1006           >> Cases_on `!i. i <= SUC (SUC n) ==> ~(x' = STRCAT "coin" (toString i))`
1007           >- (ASM_SIMP_TAC arith_ss [coin_assignment_result1]
1008               >> FULL_SIMP_TAC std_ss [IN_dc_random_states]
1009               >> Q.PAT_X_ASSUM `!x''. b ==> ~x x''` MATCH_MP_TAC
1010               >> STRIP_TAC >> REVERSE (Cases_on `i <= SUC (SUC n)`) >> ASM_REWRITE_TAC []
1011               >> FULL_SIMP_TAC arith_ss [])
1012           >> FULL_SIMP_TAC std_ss [] >> Q.PAT_X_ASSUM `i <= SUC (SUC n)` MP_TAC
1013           >> Q.SPEC_TAC (`i:num`,`i:num`) >> Induct
1014           >- (ASM_SIMP_TAC arith_ss [coin_assignment_result2, coin_assignment, T_xor, XOR_announces_def]
1015               >> `valid_coin_assignment x out p n 0` by METIS_TAC [valid_coin_assignment_result1, ZERO_LESS_EQ]
1016               >> FULL_SIMP_TAC std_ss [valid_coin_assignment, XOR_announces_def, T_xor])
1017           >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
1018           >> ASM_SIMP_TAC arith_ss [coin_assignment_result2, coin_assignment, T_xor, XOR_announces_def]
1019           >> `valid_coin_assignment x out p n (SUC i')`
1020                by (MATCH_MP_TAC valid_coin_assignment_result1 >> Q.EXISTS_TAC `SUC (SUC n)` >> FULL_SIMP_TAC arith_ss [])
1021           >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def])
1022           >> POP_ORW >> RW_TAC bool_ss [T_xor])
1023       >> DISJ2_TAC
1024       >> RW_TAC std_ss [FUN_EQ_THM]
1025       >> Cases_on `!i. i <= SUC (SUC n) ==> ~(x' = STRCAT "coin" (toString i))`
1026       >- (ASM_SIMP_TAC arith_ss [coin_assignment_result1]
1027           >> FULL_SIMP_TAC std_ss [IN_dc_random_states]
1028           >> Q.PAT_X_ASSUM `!x''. b ==> ~x x''` MATCH_MP_TAC
1029           >> STRIP_TAC >> REVERSE (Cases_on `i <= SUC (SUC n)`) >> ASM_REWRITE_TAC []
1030           >> FULL_SIMP_TAC arith_ss [])
1031       >> FULL_SIMP_TAC std_ss [] >> Q.PAT_X_ASSUM `i <= SUC (SUC n)` MP_TAC
1032       >> Q.SPEC_TAC (`i:num`,`i:num`) >> Induct
1033       >- (ASM_SIMP_TAC arith_ss [coin_assignment_result2, coin_assignment, T_xor, XOR_announces_def]
1034           >> `valid_coin_assignment x out p n 0` by METIS_TAC [valid_coin_assignment_result1, ZERO_LESS_EQ]
1035           >> FULL_SIMP_TAC std_ss [valid_coin_assignment, XOR_announces_def, T_xor])
1036       >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
1037       >> ASM_SIMP_TAC arith_ss [coin_assignment_result2, coin_assignment, T_xor, XOR_announces_def]
1038       >> `valid_coin_assignment x out p n (SUC i')`
1039           by (MATCH_MP_TAC valid_coin_assignment_result1 >> Q.EXISTS_TAC `SUC (SUC n)` >> FULL_SIMP_TAC arith_ss [])
1040       >> POP_ASSUM (ASSUME_TAC o REWRITE_RULE [valid_coin_assignment, XOR_announces_def])
1041       >> POP_ORW >> RW_TAC bool_ss [T_xor])
1042   >> STRIP_TAC
1043   >- (`x IN dc_random_states (SUC (SUC n))`
1044        by (RW_TAC std_ss [IN_dc_random_states] >> MATCH_MP_TAC coin_assignment_result1 >> METIS_TAC [])
1045       >> ASM_REWRITE_TAC []
1046       >> Suff `!i. i <= SUC (SUC n) ==> valid_coin_assignment (coin_assignment out p n T (SUC(SUC n))) out p n i`
1047       >- RW_TAC std_ss []
1048       >> Induct
1049       >- (RW_TAC arith_ss [valid_coin_assignment, XOR_announces_def, coin_assignment_result2]
1050           >> ONCE_REWRITE_TAC [coin_assignment]
1051           >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj, T_xor, xor_T]
1052           >> CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [XOR_announces_def]))
1053           >> REPEAT (POP_ASSUM MP_TAC)
1054           >> RW_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION, T_xor])
1055       >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
1056       >> ONCE_REWRITE_TAC [valid_coin_assignment] >> ASM_REWRITE_TAC []
1057       >> RW_TAC arith_ss [coin_assignment_result2]
1058       >> ONCE_REWRITE_TAC [coin_assignment]
1059       >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj, T_xor, xor_T]
1060       >> REPEAT (POP_ASSUM MP_TAC)
1061       >> RW_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION, T_xor])
1062   >> `x IN dc_random_states (SUC (SUC n))`
1063        by (RW_TAC std_ss [IN_dc_random_states] >> MATCH_MP_TAC coin_assignment_result1 >> METIS_TAC [])
1064   >> ASM_REWRITE_TAC []
1065   >> Suff `!i. i <= SUC (SUC n) ==> valid_coin_assignment (coin_assignment out p n F (SUC(SUC n))) out p n i`
1066   >- RW_TAC std_ss []
1067   >> Induct
1068   >- (RW_TAC arith_ss [valid_coin_assignment, XOR_announces_def, coin_assignment_result2]
1069       >> ONCE_REWRITE_TAC [coin_assignment]
1070       >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj, T_xor, xor_T]
1071       >> CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [XOR_announces_def]))
1072       >> REPEAT (POP_ASSUM MP_TAC)
1073       >> RW_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION, T_xor, xor_F, xor_comm])
1074   >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
1075   >> ONCE_REWRITE_TAC [valid_coin_assignment] >> ASM_REWRITE_TAC []
1076   >> RW_TAC arith_ss [coin_assignment_result2]
1077   >> ONCE_REWRITE_TAC [coin_assignment]
1078   >> FULL_SIMP_TAC arith_ss [STRCAT_toString_inj, T_xor, xor_T]
1079   >> REPEAT (POP_ASSUM MP_TAC)
1080   >> RW_TAC std_ss [dc_valid_outputs_eq_outputs, dc_valid_outputs_def, GSPECIFICATION, T_xor, xor_F, xor_comm]);
1081
1082val coin_assignment_result3 = prove
1083  (``!out p n i. i <= SUC (SUC n) ==>
1084        ~(coin_assignment out p n T i = coin_assignment out p n F i)``,
1085
1086   NTAC 3 STRIP_TAC >> Induct >- (RW_TAC bool_ss [coin_assignment, xor_F, xor_comm, T_xor, xor_def] >> METIS_TAC [])
1087   >> STRIP_TAC >> FULL_SIMP_TAC arith_ss []
1088   >> ONCE_REWRITE_TAC [coin_assignment]
1089   >> RW_TAC std_ss [FUN_EQ_THM] >> Q.EXISTS_TAC `STRCAT "coin" (toString (SUC i))`
1090   >> RW_TAC std_ss [T_xor, xor_F, xor_comm, xor_def] >> DECIDE_TAC);
1091
1092(* ------------------------------------------------------------------------- *)
1093
1094val lg_times_compute_simp_lem = prove
1095   (``!x y. x * lg (y * x) = (\x. x * lg (y * x)) x``,
1096    RW_TAC std_ss []);
1097
1098val dc_leakage_result = store_thm
1099  ("dc_leakage_result",
1100   ``!n. leakage (dc_prog_space (SUC (SUC (SUC n))) F) (dcprog (SUC (SUC (SUC n)))) = 0``,
1101   RW_TAC bool_ss [dc_prog_space_F_set_thm, unif_prog_space_leakage_computation_reduce,
1102                     dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1103                     dc_states_cross_not_empty, CARD_dc_high_states_set, CARD_dc_low_states, CARD_dc_random_states,
1104                     FINITE_CROSS, IMAGE_FINITE]
1105   >> `FINITE (IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1106      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1107       dc_random_states (SUC (SUC n))))`
1108                by RW_TAC std_ss [FINITE_CROSS, IMAGE_FINITE, dc_high_states_set_finite, dc_low_states_finite,
1109                                  dc_random_states_finite]
1110   >> `FINITE (IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1111      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1112       dc_random_states (SUC (SUC n))))`
1113                by RW_TAC std_ss [FINITE_CROSS, IMAGE_FINITE, dc_high_states_set_finite, dc_low_states_finite,
1114                                  dc_random_states_finite]
1115   >> `FINITE (IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1116      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1117       dc_random_states (SUC (SUC n))))`
1118                by RW_TAC std_ss [FINITE_CROSS, IMAGE_FINITE, dc_high_states_set_finite, dc_low_states_finite,
1119                                  dc_random_states_finite]
1120   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `(IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1121      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1122       dc_random_states (SUC (SUC n))))`) REAL_SUM_IMAGE_IN_IF]
1123   >> ONCE_REWRITE_TAC [lg_times_compute_simp_lem]
1124   >> `(\x. (if x IN IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1125                                 (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1126                                  dc_random_states (SUC (SUC n)))
1127                   then
1128                        (\(out,h,l). (\x. x * lg (1 / & (2 ** SUC (SUC (SUC n))) * x))
1129                        (SIGMA (\r. (if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then 1  else 0))
1130                                (dc_random_states (SUC (SUC n))))) x
1131                   else 0))=
1132        (\(x :bool state # bool state # bool state).
1133        (if x IN IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1134                (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1135                 dc_random_states (SUC (SUC n)))
1136         then (\x. 2 * lg ((1 :real) / & ((2 :num) ** SUC (SUC (SUC n))) * 2)) x
1137         else (0 : real)))`
1138        by (ONCE_REWRITE_TAC [FUN_EQ_THM]
1139            >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, IN_dc_high_states_set, dc_low_states_def, IN_SING]
1140            >> `FST s' = (FST (FST s'), SND (FST s'))` by RW_TAC std_ss [PAIR]
1141            >> POP_ORW
1142            >> RW_TAC std_ss []
1143            >> Suff `SIGMA (\(r :bool state). (if dcprog
1144                                (SUC (SUC (SUC (n :num)))) (FST (s' :(bool, bool, bool) prog_state),r) =
1145                                dcprog (SUC (SUC (SUC n))) s'
1146                        then (1 : real) else (0 : real))) (dc_random_states (SUC (SUC n))) = 2`
1147             >- RW_TAC std_ss []
1148             >> `(dc_random_states (SUC (SUC n))) =
1149                 {r: bool state | r IN dc_random_states (SUC (SUC n)) /\
1150                      (dcprog (SUC (SUC (SUC n)))
1151                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1152                        dcprog (SUC (SUC (SUC n))) s')} UNION
1153                 {r:bool state | r IN dc_random_states (SUC (SUC n)) /\
1154                      ~(dcprog (SUC (SUC (SUC n)))
1155                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1156                        dcprog (SUC (SUC (SUC n))) s')}`
1157                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION] >> DECIDE_TAC)
1158             >> POP_ORW
1159             >> `DISJOINT {r: bool state | r IN dc_random_states (SUC (SUC n)) /\
1160                      (dcprog (SUC (SUC (SUC n)))
1161                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1162                        dcprog (SUC (SUC (SUC n))) s')}
1163                        {r:bool state | r IN dc_random_states (SUC (SUC n)) /\
1164                      ~(dcprog (SUC (SUC (SUC n)))
1165                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1166                        dcprog (SUC (SUC (SUC n))) s')}`
1167                by (RW_TAC std_ss [DISJOINT_DEF]
1168                    >> ONCE_REWRITE_TAC [EXTENSION]
1169                    >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, DISJOINT_DEF, NOT_IN_EMPTY] >> DECIDE_TAC)
1170             >> `FINITE {r: bool state | r IN dc_random_states (SUC (SUC n)) /\
1171                      (dcprog (SUC (SUC (SUC n)))
1172                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1173                        dcprog (SUC (SUC (SUC n))) s')} /\
1174                 FINITE {r:bool state | r IN dc_random_states (SUC (SUC n)) /\
1175                      ~(dcprog (SUC (SUC (SUC n)))
1176                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1177                        dcprog (SUC (SUC (SUC n))) s')}`
1178                by (CONJ_TAC
1179                    >> (MATCH_MP_TAC o
1180                        REWRITE_RULE [dc_random_states_finite] o
1181                        Q.ISPEC `dc_random_states (SUC (SUC n))`) SUBSET_FINITE
1182                    >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
1183               >> ASM_SIMP_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION]
1184               >> (MP_TAC o Q.ISPEC `0:real` o
1185                   Q.ISPEC `(\(r :bool state). (if dcprog (SUC (SUC (SUC (n :num))))
1186                                                    (FST (s' :(bool, bool, bool) prog_state),r) =
1187                                                    dcprog (SUC (SUC (SUC n))) s'
1188                                                then (1 : real) else (0 : real)))` o
1189                   UNDISCH o Q.ISPEC `{r:bool state | r IN dc_random_states (SUC (SUC n)) /\
1190                      ~(dcprog (SUC (SUC (SUC n)))
1191                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1192                        dcprog (SUC (SUC (SUC n))) s')}`)
1193                   REAL_SUM_IMAGE_FINITE_CONST2
1194               >> SIMP_TAC std_ss [GSPECIFICATION]
1195               >> `FST s' = ((\s. s = STRCAT "pays" (toString i)),(\s. F))`
1196                by (CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [GSYM PAIR])) >> RW_TAC std_ss [])
1197               >> POP_ORW
1198               >> RW_TAC real_ss []
1199               >> POP_ASSUM (K ALL_TAC)
1200               >> (MP_TAC o Q.ISPEC `1:real` o
1201                   Q.ISPEC `(\(r :bool state). (if dcprog (SUC (SUC (SUC (n :num))))
1202                                                    (FST (s' :(bool, bool, bool) prog_state),r) =
1203                                                    dcprog (SUC (SUC (SUC n))) s'
1204                                                then (1 : real) else (0 : real)))` o
1205                   UNDISCH o Q.ISPEC `{r:bool state | r IN dc_random_states (SUC (SUC n)) /\
1206                       (dcprog (SUC (SUC (SUC n)))
1207                        (((\s. s = STRCAT "pays" (toString i)),(\s. F)),r) =
1208                        dcprog (SUC (SUC (SUC n))) s')}`)
1209                   REAL_SUM_IMAGE_FINITE_CONST2
1210               >> SIMP_TAC std_ss [GSPECIFICATION]
1211               >> `FST s' = ((\s. s = STRCAT "pays" (toString i)),(\s. F))`
1212                by (CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [GSYM PAIR])) >> RW_TAC std_ss [])
1213               >> POP_ORW
1214               >> RW_TAC real_ss []
1215               >> POP_ASSUM (K ALL_TAC)
1216               >> `dcprog (SUC (SUC (SUC n))) s' IN
1217                        IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
1218                                (dc_high_states F (SUC (SUC (SUC n))) CROSS
1219                                dc_random_states (SUC (SUC n)))`
1220                        by (RW_TAC std_ss [IN_CROSS, IN_IMAGE, dc_high_states_def, IN_dc_high_states_set]
1221                            >> `s' = ((FST (FST s'), SND (FST s')), SND s')` by RW_TAC std_ss [PAIR]
1222                            >> POP_ORW
1223                            >> Q.EXISTS_TAC `(FST(FST s'),SND s')`
1224                            >> ASM_SIMP_TAC bool_ss [] >> ASM_SIMP_TAC std_ss []
1225                            >> Q.EXISTS_TAC `i` >> RW_TAC std_ss [])
1226               >> (MP_TAC o UNDISCH_ALL o REWRITE_RULE [GSYM AND_IMP_INTRO] o Q.SPECL [`n`, `i`,
1227                                      `dcprog (SUC (SUC (SUC n)))
1228                                        (s' :(bool, bool, bool) prog_state)`])
1229                   valid_coin_set_eq_valid_coin_assignment
1230               >> RW_TAC std_ss [valid_coin_assignment_eq_2_element_set, coin_assignment_set]
1231               >> RW_TAC set_ss [coin_assignment, FUN_EQ_THM, T_xor, xor_T, xor_F]
1232               >> METIS_TAC [])
1233   >> `!x. (\x.
1234           (\(out,h,l).
1235              (\x. x * lg (1 / &(2 ** SUC (SUC (SUC n))) * x))
1236                (SIGMA
1237                   (\r.
1238                      if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then
1239                        1
1240                      else
1241                        0) (dc_random_states (SUC (SUC n))))) x) x =
1242        (\(out,h,l).
1243              (\x. x * lg (1 / &(2 ** SUC (SUC (SUC n))) * x))
1244                (SIGMA
1245                   (\r.
1246                      if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then
1247                        1
1248                      else
1249                        0) (dc_random_states (SUC (SUC n))))) x` by METIS_TAC []
1250   >> POP_ORW
1251   >> POP_ORW
1252   >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_IN_IF, dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1253                     FINITE_CROSS, IMAGE_FINITE]
1254   >> (MP_TAC o Q.ISPEC `(IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1255      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1256       dc_random_states (SUC (SUC n))))`) REAL_SUM_IMAGE_FINITE_CONST
1257   >> RW_TAC std_ss [dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1258                     FINITE_CROSS, IMAGE_FINITE]
1259   >> POP_ASSUM (MP_TAC o Q.SPECL [`(\x. 2 * lg (1 / & (2 ** SUC (SUC (SUC n))) * 2))`,
1260                                   `2 * lg (1 / & (2 ** SUC (SUC (SUC n))) * 2)`])
1261   >> RW_TAC real_ss []
1262   >> POP_ASSUM (K ALL_TAC)
1263   >> RW_TAC bool_ss [dc_prog_space_F_set_thm, unif_prog_space_leakage_computation_reduce,
1264                     dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1265                     dc_states_cross_not_empty, CARD_dc_high_states_set, CARD_dc_low_states, CARD_dc_random_states,
1266                     FINITE_CROSS, IMAGE_FINITE]
1267   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `(IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1268      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1269       dc_random_states (SUC (SUC n))))`) REAL_SUM_IMAGE_IN_IF]
1270   >> ONCE_REWRITE_TAC [lg_times_compute_simp_lem]
1271   >> `(\x. (if x IN IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1272                        (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1273                        dc_random_states (SUC (SUC n)))
1274             then (\(out,l). (\x. x * lg (1 / & (SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) * x))
1275                (SIGMA (\(h,r). (if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then 1 else 0))
1276                 (dc_high_states_set (SUC (SUC n)) CROSS
1277                  dc_random_states (SUC (SUC n))))) x
1278             else 0)) =
1279        (\x. (if x IN IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1280                        (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1281                        dc_random_states (SUC (SUC n)))
1282             then (\x. (&(2 * (SUC(SUC(SUC n))))) * lg (1 / & (SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1283                        (&(2 * (SUC(SUC(SUC n))))))) x
1284             else 0))`
1285        by (ONCE_REWRITE_TAC [FUN_EQ_THM]
1286            >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, IN_dc_high_states_set, dc_low_states_def, IN_SING]
1287            >> `FST s' = (FST (FST s'), SND (FST s'))` by RW_TAC std_ss [PAIR]
1288            >> POP_ORW
1289            >> RW_TAC std_ss []
1290            >> Suff `SIGMA (\(h,r). (if dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1291                                        dcprog (SUC (SUC (SUC n))) s'
1292                                     then 1 else 0))
1293                        (dc_high_states_set (SUC (SUC n)) CROSS
1294                        dc_random_states (SUC (SUC n))) = & (2 * SUC (SUC (SUC n)))`
1295            >- RW_TAC real_ss []
1296            >> `(dc_high_states_set (SUC (SUC n)) CROSS dc_random_states (SUC (SUC n))) =
1297                 {(h:bool state, r: bool state) |
1298                        h IN dc_high_states_set (SUC (SUC n)) /\
1299                        r IN dc_random_states (SUC (SUC n)) /\
1300                      (dcprog (SUC (SUC (SUC n)))
1301                        ((h,(\s. F)),r) =
1302                        dcprog (SUC (SUC (SUC n))) s')} UNION
1303                 {(h:bool state, r: bool state) |
1304                        h IN dc_high_states_set (SUC (SUC n)) /\
1305                        r IN dc_random_states (SUC (SUC n)) /\
1306                      ~(dcprog (SUC (SUC (SUC n)))
1307                        ((h,(\s. F)),r) =
1308                        dcprog (SUC (SUC (SUC n))) s')}`
1309                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
1310                    >> EQ_TAC
1311                    >- (RW_TAC std_ss []
1312                        >> Cases_on `dcprog (SUC (SUC (SUC n))) ((FST x,(\s. F)),SND x) =
1313                                     dcprog (SUC (SUC (SUC n))) s'`
1314                        >- (DISJ1_TAC
1315                            >> Q.EXISTS_TAC `(FST x,SND x)`
1316                            >> RW_TAC std_ss [PAIR_EQ])
1317                        >> DISJ2_TAC
1318                        >> Q.EXISTS_TAC `(FST x,SND x)`
1319                        >> RW_TAC std_ss [PAIR_EQ])
1320                    >> RW_TAC std_ss []
1321                    >> POP_ASSUM MP_TAC
1322                    >> (MP_TAC o Q.ISPEC `x': bool state # bool state`) pair_CASES
1323                    >> RW_TAC bool_ss []
1324                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss [])
1325             >> POP_ORW
1326             >> `DISJOINT {(h:bool state, r: bool state) |
1327                        h IN dc_high_states_set (SUC (SUC n)) /\
1328                        r IN dc_random_states (SUC (SUC n)) /\
1329                      (dcprog (SUC (SUC (SUC n)))
1330                        ((h,(\s. F)),r) =
1331                        dcprog (SUC (SUC (SUC n))) s')}
1332                 {(h:bool state, r: bool state) |
1333                        h IN dc_high_states_set (SUC (SUC n)) /\
1334                        r IN dc_random_states (SUC (SUC n)) /\
1335                      ~(dcprog (SUC (SUC (SUC n)))
1336                        ((h,(\s. F)),r) =
1337                        dcprog (SUC (SUC (SUC n))) s')}`
1338                by (RW_TAC std_ss [DISJOINT_DEF]
1339                    >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, DISJOINT_DEF, NOT_IN_EMPTY]
1340                    >> Cases_on `(FST x) IN dc_high_states_set (SUC (SUC n)) /\
1341                                 (SND x) IN dc_random_states (SUC (SUC n)) /\
1342                                 (dcprog (SUC (SUC (SUC n))) ((FST x,(\s. F)),SND x) =
1343                                  dcprog (SUC (SUC (SUC n))) s')`
1344                    >- (DISJ2_TAC
1345                        >> RW_TAC std_ss []
1346                        >> (MP_TAC o Q.ISPEC `x': bool state # bool state`) pair_CASES
1347                        >> RW_TAC bool_ss [] >> RW_TAC std_ss []
1348                        >> REVERSE (Cases_on `(x = (q,r))`) >> RW_TAC std_ss []
1349                        >> FULL_SIMP_TAC std_ss [FST,SND])
1350                    >> DISJ1_TAC
1351                    >> RW_TAC std_ss []
1352                    >> (MP_TAC o Q.ISPEC `x': bool state # bool state`) pair_CASES
1353                    >> RW_TAC bool_ss [] >> RW_TAC std_ss []
1354                    >> REVERSE (Cases_on `(x = (q,r))`) >> RW_TAC std_ss []
1355                    >> FULL_SIMP_TAC std_ss [FST,SND])
1356             >> `FINITE {(h:bool state, r: bool state) |
1357                        h IN dc_high_states_set (SUC (SUC n)) /\
1358                        r IN dc_random_states (SUC (SUC n)) /\
1359                      (dcprog (SUC (SUC (SUC n)))
1360                        ((h,(\s. F)),r) =
1361                        dcprog (SUC (SUC (SUC n))) s')} /\
1362                 FINITE {(h:bool state, r: bool state) |
1363                        h IN dc_high_states_set (SUC (SUC n)) /\
1364                        r IN dc_random_states (SUC (SUC n)) /\
1365                      ~(dcprog (SUC (SUC (SUC n)))
1366                        ((h,(\s. F)),r) =
1367                        dcprog (SUC (SUC (SUC n))) s')}`
1368                by (CONJ_TAC
1369                    >> (MATCH_MP_TAC o
1370                        CONV_RULE (SIMP_CONV std_ss [dc_random_states_finite, dc_high_states_set_finite, FINITE_CROSS]) o
1371                        Q.ISPEC `(dc_high_states_set (SUC(SUC n))) CROSS
1372                                 (dc_random_states (SUC (SUC n)))`) SUBSET_FINITE
1373                    >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
1374                    >> POP_ASSUM MP_TAC
1375                    >> (MP_TAC o Q.ISPEC `x': bool state # bool state`) pair_CASES
1376                    >> RW_TAC bool_ss []
1377                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss [])
1378               >> ASM_SIMP_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION]
1379               >> `(!y. y IN {(h,r) | h IN dc_high_states_set (SUC (SUC n)) /\
1380                                       r IN dc_random_states (SUC (SUC n)) /\
1381                                       ~(dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1382                                         dcprog (SUC (SUC (SUC n))) s')} ==>
1383                        ((\(h,r). (if dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1384                                      dcprog (SUC (SUC (SUC n))) s'
1385                                   then 1 else 0)) y = 0))`
1386                        by (RW_TAC std_ss [GSPECIFICATION]
1387                            >> (MP_TAC o Q.ISPEC `x: bool state # bool state`) pair_CASES
1388                            >> RW_TAC bool_ss [] >> FULL_SIMP_TAC std_ss [])
1389               >> ((fn thm => ONCE_REWRITE_TAC [thm]) o UNDISCH o Q.ISPEC `0:real` o
1390                   Q.ISPEC `(\(h:bool state, r :bool state). (if dcprog (SUC (SUC (SUC (n :num))))
1391                                                    ((h, (\s:string. F)) ,r) =
1392                                                    dcprog (SUC (SUC (SUC n))) s'
1393                                                then (1 : real) else (0 : real)))` o
1394                   UNDISCH o Q.ISPEC `{(h:bool state, r: bool state) |
1395                        h IN dc_high_states_set (SUC (SUC n)) /\
1396                        r IN dc_random_states (SUC (SUC n)) /\
1397                      ~(dcprog (SUC (SUC (SUC n)))
1398                        ((h,(\s:string. F)),r) =
1399                        dcprog (SUC (SUC (SUC n))) s')}`)
1400                   REAL_SUM_IMAGE_FINITE_CONST2
1401               >> POP_ASSUM (K ALL_TAC) >> RW_TAC real_ss []
1402               >> `(!y. y IN {(h,r) | h IN dc_high_states_set (SUC (SUC n)) /\
1403                                       r IN dc_random_states (SUC (SUC n)) /\
1404                                       (dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1405                                         dcprog (SUC (SUC (SUC n))) s')} ==>
1406                        ((\(h,r). (if dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1407                                      dcprog (SUC (SUC (SUC n))) s'
1408                                   then 1 else 0)) y = 1))`
1409                        by (RW_TAC std_ss [GSPECIFICATION]
1410                            >> (MP_TAC o Q.ISPEC `x: bool state # bool state`) pair_CASES
1411                            >> RW_TAC bool_ss [] >> FULL_SIMP_TAC std_ss [])
1412               >> ((fn thm => ONCE_REWRITE_TAC [thm]) o UNDISCH o Q.ISPEC `1:real` o
1413                   Q.ISPEC `(\(h:bool state, r :bool state). (if dcprog (SUC (SUC (SUC (n :num))))
1414                                                    ((h, (\s:string. F)) ,r) =
1415                                                    dcprog (SUC (SUC (SUC n))) s'
1416                                                then (1 : real) else (0 : real)))` o
1417                   UNDISCH o Q.ISPEC `{(h:bool state, r: bool state) |
1418                        h IN dc_high_states_set (SUC (SUC n)) /\
1419                        r IN dc_random_states (SUC (SUC n)) /\
1420                      (dcprog (SUC (SUC (SUC n)))
1421                        ((h,(\s:string. F)),r) =
1422                        dcprog (SUC (SUC (SUC n))) s')}`)
1423                   REAL_SUM_IMAGE_FINITE_CONST2
1424               >> POP_ASSUM (K ALL_TAC) >> RW_TAC real_ss []
1425               >> `(dcprog (SUC (SUC (SUC n))) s') IN
1426                                IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
1427                                        (dc_high_states F (SUC (SUC (SUC n))) CROSS
1428                                         dc_random_states (SUC (SUC n)))`
1429                                by (RW_TAC std_ss [IN_IMAGE, IN_CROSS, dc_high_states_def]
1430                                            >> Q.EXISTS_TAC `(FST(FST s'),SND s')`
1431                                            >> RW_TAC std_ss [IN_dc_high_states_set]
1432                                            >- (`s' = ((FST (FST s'), SND (FST s')), SND s')` by RW_TAC std_ss [PAIR]
1433                                                >> POP_ORW
1434                                                >> RW_TAC bool_ss [FST,SND])
1435                                            >> Q.EXISTS_TAC `i` >> RW_TAC std_ss [])
1436               >> `!h r.  (h IN dc_high_states_set (SUC (SUC n)) /\
1437                           r IN dc_random_states (SUC (SUC n))  /\
1438                         (dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r) =
1439                          dcprog (SUC (SUC (SUC n))) s')) =
1440                        (?p. p <= SUC (SUC n) /\ (h = (\s. s = STRCAT "pays" (toString p))) /\
1441                             r IN {r | r IN dc_random_states (SUC (SUC n))  /\
1442                          (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString p)),(\s. F)),r) =
1443                           dcprog (SUC (SUC (SUC n))) s')})`
1444                by (RW_TAC std_ss [IN_dc_high_states_set, GSPECIFICATION] >> EQ_TAC
1445                    >- (STRIP_TAC >> `(h = (\s. s = STRCAT "pays" (toString i')))` by METIS_TAC []
1446                        >> Q.EXISTS_TAC `i'` >> RW_TAC std_ss [])
1447                    >> RW_TAC std_ss [] >- (Q.EXISTS_TAC `p` >> RW_TAC std_ss []) >> ASM_REWRITE_TAC [])
1448               >> POP_ORW
1449               >> `!h r. (?p. p <= SUC (SUC n) /\ (h = (\s. s = STRCAT "pays" (toString p))) /\
1450                             r IN {r | r IN dc_random_states (SUC (SUC n))  /\
1451                          (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString p)),(\s. F)),r) =
1452                           dcprog (SUC (SUC (SUC n))) s')}) =
1453                          (?p.  p <= SUC (SUC n) /\
1454                                (h = (\s. s = STRCAT "pays" (toString p))) /\
1455                                r IN coin_assignment_set (dcprog (SUC (SUC (SUC n))) s') p n)`
1456                        by (REPEAT STRIP_TAC
1457                            >> EQ_TAC
1458                            >- (STRIP_TAC >> Q.EXISTS_TAC `p` >> ASM_REWRITE_TAC []
1459                                >> `{r | r IN dc_random_states (SUC (SUC n))  /\
1460                                         (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString p)),(\s. F)),r) =
1461                                          dcprog (SUC (SUC (SUC n))) s')} =
1462                                    {r | r IN dc_random_states (SUC (SUC n)) /\
1463                                         valid_coin_assignment r (dcprog (SUC (SUC (SUC n))) s') p n (SUC (SUC n))}`
1464                                        by (MATCH_MP_TAC valid_coin_set_eq_valid_coin_assignment
1465                                            >> ASM_REWRITE_TAC [])
1466                                >> FULL_SIMP_TAC std_ss []
1467                                >> Suff `{r | r IN dc_random_states (SUC (SUC n)) /\
1468                                                  valid_coin_assignment r (dcprog (SUC (SUC (SUC n))) s') p n (SUC (SUC n))} =
1469                                             coin_assignment_set (dcprog (SUC (SUC (SUC n))) s') p n`
1470                                >- (STRIP_TAC >> FULL_SIMP_TAC std_ss [])
1471                                >> MATCH_MP_TAC valid_coin_assignment_eq_2_element_set
1472                                >> ASM_REWRITE_TAC [])
1473                            >> STRIP_TAC >> Q.EXISTS_TAC `p` >> RW_TAC std_ss []
1474                            >> Suff `{r | r IN dc_random_states (SUC (SUC n)) /\
1475                                              (dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString p)),(\s. F)),r) =
1476                                              (dcprog (SUC (SUC (SUC n))) s'))} =
1477                                                 coin_assignment_set (dcprog (SUC (SUC (SUC n))) s') p n`
1478                            >- (STRIP_TAC >> FULL_SIMP_TAC std_ss [])
1479                            >> ASM_SIMP_TAC std_ss [valid_coin_set_eq_valid_coin_assignment,
1480                                                    valid_coin_assignment_eq_2_element_set])
1481               >> POP_ORW
1482               >> RW_TAC std_ss [coin_assignment_set, IN_INSERT, NOT_IN_EMPTY]
1483               >> `{(h,r) | ?p. p <= SUC (SUC n) /\
1484                                (h = (\s. s = STRCAT "pays" (toString p))) /\
1485                                ((r = coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n))) \/
1486                                (r = coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n))))} =
1487                   {(h,r) | ?p. p <= SUC (SUC n) /\
1488                                (h = (\s. s = STRCAT "pays" (toString p))) /\
1489                                ((r = coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n))))} UNION
1490                   {(h,r) | ?p. p <= SUC (SUC n) /\
1491                                (h = (\s. s = STRCAT "pays" (toString p))) /\
1492                                ((r = coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n))))}`
1493                 by (RW_TAC std_ss [EXTENSION, GSPECIFICATION, IN_UNION] >> RW_TAC std_ss [GSYM EXTENSION]
1494                     >> EQ_TAC
1495                     >- (STRIP_TAC >> (ASSUME_TAC o Q.ISPEC `x':bool state # bool state`) pair_CASES
1496                         >> FULL_SIMP_TAC std_ss [] >> FULL_SIMP_TAC std_ss [PAIR_EQ]
1497                         >- (DISJ1_TAC >> Q.EXISTS_TAC `(q,r)` >> RW_TAC std_ss []
1498                             >> `(?p'. p' <= SUC (SUC n) /\
1499                                        ((\s. s = STRCAT "pays" (toString p)) =
1500                                         (\s. s = STRCAT "pays" (toString p'))) /\
1501                                (coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n)) =
1502                                 coin_assignment (dcprog (SUC (SUC (SUC n))) s') p' n T (SUC (SUC n)))) = T`
1503                                by (RW_TAC std_ss [] >> Q.EXISTS_TAC `p` >> RW_TAC std_ss [])
1504                             >> POP_ORW
1505                             >> RW_TAC std_ss [] >> Q.EXISTS_TAC `p`
1506                             >> RW_TAC std_ss [])
1507                         >> DISJ2_TAC >> Q.EXISTS_TAC `(q,r)` >> RW_TAC std_ss []
1508                         >> `(?p'. p' <= SUC (SUC n) /\
1509                                ((\s. s = STRCAT "pays" (toString p)) =
1510                                 (\s. s = STRCAT "pays" (toString p'))) /\
1511                                (coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n)) =
1512                                 coin_assignment (dcprog (SUC (SUC (SUC n))) s') p' n F (SUC (SUC n)))) = T`
1513                                by (RW_TAC std_ss [] >> Q.EXISTS_TAC `p` >> RW_TAC std_ss [])
1514                         >> POP_ORW
1515                         >> RW_TAC std_ss [] >> Q.EXISTS_TAC `p`
1516                         >> RW_TAC std_ss [])
1517                     >> STRIP_TAC
1518                     >- ((ASSUME_TAC o Q.ISPEC `x':bool state # bool state`) pair_CASES
1519                         >> FULL_SIMP_TAC std_ss [] >> FULL_SIMP_TAC std_ss [PAIR_EQ]
1520                         >> Q.EXISTS_TAC `(q,r)` >> RW_TAC std_ss []
1521                         >> `(?p'. p' <= SUC (SUC n) /\
1522                        ((\s. s = STRCAT "pays" (toString p)) =
1523                        (\s. s = STRCAT "pays" (toString p'))) /\
1524                        (coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n)) =
1525                         coin_assignment (dcprog (SUC (SUC (SUC n))) s') p' n T (SUC (SUC n)))) = T`
1526                        by (RW_TAC std_ss [] >> Q.EXISTS_TAC `p` >> RW_TAC std_ss [])
1527                         >> POP_ORW
1528                         >> RW_TAC std_ss [] >> Q.EXISTS_TAC `p`
1529                         >> RW_TAC std_ss [])
1530                     >> (ASSUME_TAC o Q.ISPEC `x':bool state # bool state`) pair_CASES
1531                     >> FULL_SIMP_TAC std_ss [] >> FULL_SIMP_TAC std_ss [PAIR_EQ]
1532                     >> Q.EXISTS_TAC `(q,r)` >> RW_TAC std_ss []
1533                     >> `(?p'. p' <= SUC (SUC n) /\
1534                        ((\s. s = STRCAT "pays" (toString p)) =
1535                        (\s. s = STRCAT "pays" (toString p'))) /\
1536                        (coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n)) =
1537                         coin_assignment (dcprog (SUC (SUC (SUC n))) s') p' n F (SUC (SUC n)))) = T`
1538                        by (RW_TAC std_ss [] >> Q.EXISTS_TAC `p` >> RW_TAC std_ss [])
1539                     >> POP_ORW
1540                     >> RW_TAC std_ss [] >> Q.EXISTS_TAC `p`
1541                     >> RW_TAC std_ss [])
1542               >> POP_ORW
1543               >> `!b. {(h,r) | ?p. p <= SUC (SUC n) /\
1544                                (h = (\s. s = STRCAT "pays" (toString p))) /\
1545                                (r = coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n b (SUC (SUC n)))} =
1546                        IMAGE (\p:num. ((\s:string. s = STRCAT "pays" (toString p)),
1547                        coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n b (SUC (SUC n))))
1548                        (pred_set$count (SUC(SUC(SUC n))))`
1549                by (RW_TAC std_ss [EXTENSION, GSPECIFICATION, IN_IMAGE, pred_setTheory.IN_COUNT]
1550                    >> RW_TAC std_ss [GSYM EXTENSION]
1551                    >> EQ_TAC >- (STRIP_TAC >> (ASSUME_TAC o Q.ISPEC `x':bool state # bool state`) pair_CASES
1552                                  >> FULL_SIMP_TAC std_ss [] >> FULL_SIMP_TAC std_ss [PAIR_EQ]
1553                                  >> Q.EXISTS_TAC `p` >> FULL_SIMP_TAC arith_ss [])
1554                    >> STRIP_TAC >> Q.EXISTS_TAC `x` >> RW_TAC std_ss [PAIR_EQ] >> Q.EXISTS_TAC `p`
1555                    >> FULL_SIMP_TAC arith_ss [])
1556               >> POP_ORW >> RW_TAC std_ss [IMAGE_UNION]
1557               >> (ASSUME_TAC o Q.ISPEC `IMAGE (\p. (((\s. s = STRCAT "pays" (toString p))),
1558                        coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n))))
1559                                (pred_set$count (SUC (SUC (SUC n))))`) CARD_UNION
1560               >> FULL_SIMP_TAC std_ss [IMAGE_FINITE, pred_setTheory.FINITE_COUNT]
1561               >> POP_ASSUM (ASSUME_TAC o Q.ISPEC `IMAGE (\p. (((\s. s = STRCAT "pays" (toString p))),
1562                        coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n))))
1563                                (pred_set$count (SUC (SUC (SUC n))))`)
1564               >> FULL_SIMP_TAC std_ss [IMAGE_FINITE, pred_setTheory.FINITE_COUNT]
1565               >> `(IMAGE (\p.
1566                  (((\s. s = STRCAT "pays" (toString p))),
1567                   coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n))))
1568                        (pred_set$count (SUC (SUC (SUC n)))) INTER
1569                    IMAGE (\p.
1570                  (((\s. s = STRCAT "pays" (toString p))),
1571                   coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n F (SUC (SUC n))))
1572                        (pred_set$count (SUC (SUC (SUC n))))) = {}`
1573                by (RW_TAC std_ss [EXTENSION, INTER_DEF, IN_IMAGE, NOT_IN_EMPTY, pred_setTheory.IN_COUNT, GSPECIFICATION]
1574                    >> Cases_on `(!p.  ~(x = (((\s. s = STRCAT "pays" (toString p))),
1575                                coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n T (SUC (SUC n)))) \/
1576                                ~(p < SUC (SUC (SUC n))))`
1577                    >> ASM_REWRITE_TAC [] >> FULL_SIMP_TAC std_ss []
1578                    >> RW_TAC std_ss [] >> REVERSE (Cases_on `p' < SUC (SUC (SUC n))`) >> ASM_REWRITE_TAC []
1579                    >> REVERSE (Cases_on `((\s. s = STRCAT "pays" (toString p)) =
1580                                  (\s. s = STRCAT "pays" (toString p')))`) >> ASM_REWRITE_TAC []
1581                    >> `p' = p` by METIS_TAC [STRCAT_toString_inj]
1582                    >> ASM_SIMP_TAC arith_ss [coin_assignment_result3, LESS_EQ_REFL])
1583               >> FULL_SIMP_TAC std_ss [CARD_EMPTY] >> POP_ASSUM (K ALL_TAC) >> POP_ASSUM (K ALL_TAC)
1584               >> `!b. CARD (IMAGE (\p. (((\s. s = STRCAT "pays" (toString p))),
1585                        coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n b (SUC (SUC n))))
1586                        (pred_set$count (SUC (SUC (SUC n))))) = CARD (pred_set$count (SUC (SUC (SUC n))))`
1587                by (STRIP_TAC
1588                    >> (MATCH_MP_TAC o REWRITE_RULE [pred_setTheory.FINITE_COUNT] o
1589                        Q.ISPECL [`(\p. (((\s. s = STRCAT "pays" (toString p))),
1590                                        coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n b
1591                                        (SUC (SUC n))))`,
1592                                  `(pred_set$count (SUC (SUC (SUC (n:num)))))`,
1593                                  `IMAGE (\p. (((\s. s = STRCAT "pays" (toString p))),
1594                                                coin_assignment (dcprog (SUC (SUC (SUC n))) s') p n b
1595                                                (SUC (SUC n))))
1596                                         (pred_set$count (SUC (SUC (SUC (n:num)))))`]) CARD_IMAGE
1597                     >> RW_TAC std_ss [INJ_DEF, pred_setTheory.IN_COUNT, PAIR_EQ, IN_IMAGE]
1598                     >> METIS_TAC [STRCAT_toString_inj])
1599               >> POP_ORW >> RW_TAC arith_ss [pred_setTheory.CARD_COUNT])
1600   >> `!x. (\x.
1601           (\(out,l).
1602              (\x.
1603                 x *
1604                 lg
1605                   (1 / &(SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1606                    x))
1607                (SIGMA
1608                   (\(h,r).
1609                      if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then
1610                        1
1611                      else
1612                        0)
1613                   (dc_high_states_set (SUC (SUC n)) CROSS
1614                    dc_random_states (SUC (SUC n))))) x) x =
1615        (\(out,l).
1616              (\x.
1617                 x *
1618                 lg
1619                   (1 / &(SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1620                    x))
1621                (SIGMA
1622                   (\(h,r).
1623                      if dcprog (SUC (SUC (SUC n))) ((h,l),r) = out then
1624                        1
1625                      else
1626                        0)
1627                   (dc_high_states_set (SUC (SUC n)) CROSS
1628                    dc_random_states (SUC (SUC n))))) x`
1629        by METIS_TAC []
1630   >> POP_ORW
1631   >> POP_ORW
1632   >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_IN_IF, dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1633                     FINITE_CROSS, IMAGE_FINITE]
1634   >> (MP_TAC o Q.ISPEC `(IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1635      (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1636       dc_random_states (SUC (SUC n))))`) REAL_SUM_IMAGE_FINITE_CONST
1637   >> RW_TAC std_ss [dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1638                     FINITE_CROSS, IMAGE_FINITE]
1639   >> POP_ASSUM (MP_TAC o Q.SPECL [`(\x. & (2 * SUC (SUC (SUC n))) *
1640                                        lg (1 / & (SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1641                                        & (2 * SUC (SUC (SUC n)))))`,
1642                                   `& (2 * SUC (SUC (SUC n))) *
1643                                        lg (1 / & (SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1644                                        & (2 * SUC (SUC (SUC n))))`])
1645   >> RW_TAC real_ss []
1646   >> POP_ASSUM (K ALL_TAC)
1647   >> Suff `& (CARD
1648      (IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1649         (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1650          dc_random_states (SUC (SUC n))))) * (2 * lg (1 / & (2 ** SUC (SUC (SUC n))) * 2)) =
1651        & (CARD
1652      (IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1653         (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1654          dc_random_states (SUC (SUC n))))) * (& (2 * SUC (SUC (SUC n))) *
1655        lg (1 / & (SUC (SUC (SUC n)) * 2 ** SUC (SUC (SUC n))) *
1656        & (2 * SUC (SUC (SUC n)))))`
1657   >- RW_TAC real_ss []
1658   >> `(IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,SND (FST s)))
1659        (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1660         dc_random_states (SUC (SUC n)))) =
1661        (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
1662          (dc_high_states F (SUC (SUC (SUC n))) CROSS
1663           dc_random_states (SUC (SUC n)))) CROSS dc_low_states`
1664        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_CROSS, IN_IMAGE, dc_low_states_def, IN_SING]
1665            >> EQ_TAC
1666            >- (RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND]
1667                >> Q.EXISTS_TAC `(FST (FST s'), SND s')`
1668                >> RW_TAC bool_ss [FST,SND, dc_high_states_def]
1669                >> RW_TAC std_ss []
1670                >> `s' = ((FST (FST s'), SND (FST s')), SND s')` by RW_TAC std_ss [PAIR]
1671                >> POP_ORW >> RW_TAC bool_ss [FST, SND])
1672            >> RW_TAC std_ss [dc_high_states_def] >> RW_TAC std_ss [FST,SND]
1673            >> Q.EXISTS_TAC `((FST x', (\s:string.F)),SND x')`
1674            >> RW_TAC std_ss [] >> CONV_TAC (RATOR_CONV (ONCE_REWRITE_CONV [GSYM PAIR]))
1675            >> ASM_SIMP_TAC bool_ss [PAIR_EQ]
1676            >> `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
1677            >> POP_ORW >> RW_TAC bool_ss [FST,SND] >> RW_TAC std_ss [])
1678    >> POP_ORW
1679    >> `IMAGE (\s. (dcprog (SUC (SUC (SUC n))) s,FST s))
1680        (dc_high_states_set (SUC (SUC n)) CROSS dc_low_states CROSS
1681         dc_random_states (SUC (SUC n))) =
1682             (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s:string. F)),r))
1683          (dc_high_states F (SUC (SUC (SUC n))) CROSS
1684           dc_random_states (SUC (SUC n)))) CROSS ((dc_high_states_set (SUC (SUC n)) CROSS dc_low_states))`
1685        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [dc_low_states_def, dc_high_states_def, IN_CROSS, IN_SING, IN_IMAGE]
1686            >> EQ_TAC
1687            >- (RW_TAC std_ss [] >> FULL_SIMP_TAC std_ss [FST, SND]
1688                >> Q.EXISTS_TAC `(FST (FST s'),SND s')`
1689                >> RW_TAC std_ss [FST, SND, dc_high_states_def]
1690                >- (POP_ASSUM (K ALL_TAC) >> POP_ASSUM (MP_TAC o GSYM) >> RW_TAC bool_ss [PAIR])
1691                >> FULL_SIMP_TAC std_ss [IN_dc_random_states]
1692                >> RW_TAC std_ss []
1693                >> Q.PAT_X_ASSUM `!x. P ==> ~SND s' x` MATCH_MP_TAC
1694                >> METIS_TAC [DECIDE ``!(i:num) (n:num). i <= n ==> i <= SUC n``])
1695            >> RW_TAC std_ss [IN_dc_high_states_set]
1696            >> REPEAT (POP_ASSUM MP_TAC)
1697            >> ONCE_REWRITE_TAC [(DECIDE ``(x' :bool state # bool state) = (FST x', SND x')``)]
1698            >> RW_TAC std_ss [FST, SND]
1699            >> `x = (FST x, SND x)` by RW_TAC std_ss [PAIR]
1700            >> POP_ORW
1701            >> RW_TAC bool_ss [PAIR_EQ]
1702            >> Q.ABBREV_TAC `r = coin_assignment (dcprog (SUC (SUC (SUC n))) ((FST x',(\s:string. F)),SND x'))
1703                                                i' n T (SUC(SUC n))`
1704            >> Q.ABBREV_TAC `out = dcprog (SUC (SUC (SUC n))) (((\s. s = STRCAT "pays" (toString i)),(\s:string. F)),SND x')`
1705            >> `out IN (IMAGE (\(h,r). dcprog (SUC (SUC (SUC n))) ((h,(\s. F)),r))
1706                             (dc_high_states F (SUC (SUC (SUC n))) CROSS dc_random_states (SUC (SUC n))))`
1707                by (Q.UNABBREV_TAC `out`
1708                    >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, dc_high_states_def, IN_dc_high_states_set, dc_low_states_def, IN_SING]
1709                    >> Q.EXISTS_TAC `((\s. s = STRCAT "pays" (toString i)),SND x')`
1710                    >> RW_TAC std_ss [FST,SND] >> Q.EXISTS_TAC `i` >> RW_TAC std_ss [])
1711            >> `r IN {r:bool state| r IN (dc_random_states (SUC (SUC n))) /\
1712                                    (dcprog (SUC(SUC(SUC n))) (((\s. s = STRCAT "pays" (toString i')), (\s:string. F)), r)=out)}`
1713                by (RW_TAC std_ss [valid_coin_set_eq_valid_coin_assignment, valid_coin_assignment_eq_2_element_set,
1714                                   coin_assignment_set]
1715                    >> Q.UNABBREV_TAC `r`
1716                    >> ASM_SIMP_TAC set_ss [])
1717            >> `SND x = (FST(SND x),SND(SND x))` by RW_TAC std_ss [PAIR]
1718            >> POP_ORW
1719            >> Q.EXISTS_TAC `(((\s. s = STRCAT "pays" (toString i')),(\s:string.F)),r)`
1720            >> RW_TAC bool_ss [FST,SND] >> FULL_SIMP_TAC std_ss [GSPECIFICATION]
1721            >> Q.EXISTS_TAC `i'` >> ASM_REWRITE_TAC [])
1722   >> RW_TAC std_ss [CARD_CROSS, dc_high_states_set_finite, dc_random_states_finite, dc_low_states_finite,
1723                      FINITE_CROSS, IMAGE_FINITE, CARD_dc_low_states, dc_high_states_def, CARD_dc_high_states_set,
1724                      REWRITE_RULE [dc_high_states_def] CARD_dc_valid_outputs]
1725   >> `~(2 ** SUC (SUC (SUC n)) = 0)` by RW_TAC arith_ss [EXP_EQ_0]
1726   >> `~(2 ** (SUC (SUC n)) = 0)` by RW_TAC arith_ss [EXP_EQ_0]
1727   >> `~(2 pow (SUC (SUC n)) = 0)`
1728                by (ONCE_REWRITE_TAC [EQ_SYM_EQ] >> MATCH_MP_TAC REAL_LT_IMP_NE
1729                    >> MATCH_MP_TAC REAL_LET_TRANS >> Q.EXISTS_TAC `&((SUC(SUC n)))` >> RW_TAC real_ss [POW_2_LT])
1730   >> `~(2 pow (SUC(SUC (SUC n))) = 0)`
1731                by (ONCE_REWRITE_TAC [EQ_SYM_EQ] >> MATCH_MP_TAC REAL_LT_IMP_NE
1732                    >> MATCH_MP_TAC REAL_LET_TRANS >> Q.EXISTS_TAC `&((SUC(SUC(SUC n))))` >> RW_TAC real_ss [POW_2_LT])
1733   >> RW_TAC std_ss [GSYM REAL_MUL]
1734   >> RW_TAC real_ss [GSYM REAL_INV_1OVER, REAL_INV_MUL, real_div, REAL_INV_INV, GSYM REAL_OF_NUM_POW,
1735                           lg_mul, REAL_INV_EQ_0, REAL_INV_MUL]
1736   >> RW_TAC std_ss [GSYM REAL_MUL]
1737   >> RW_TAC std_ss [GSYM REAL_OF_NUM_POW]
1738   >> Suff `lg (inv (& (SUC (SUC (SUC n)))) * inv (2 pow SUC (SUC (SUC n))) *
1739                 (2 * & (SUC (SUC (SUC n))))) = lg (inv (2 pow SUC (SUC (SUC n))) * 2)`
1740   >- (RW_TAC std_ss [] >> REAL_ARITH_TAC)
1741   >> `inv (& (SUC (SUC (SUC n)))) * inv (2 pow SUC (SUC (SUC n))) * (2 * & (SUC (SUC (SUC n)))) =
1742            (& (SUC (SUC (SUC n))) * inv (& (SUC (SUC (SUC n))))) * inv (2 pow SUC (SUC (SUC n))) * 2`
1743                by REAL_ARITH_TAC
1744   >> POP_ORW >> RW_TAC real_ss [REAL_MUL_RINV]);
1745
1746(* ************************************************************************* *)
1747(* Case Study: The Dining Cryptographers - From the view of a cryptographer  *)
1748(* for the sake of simplicity, assume crypto 0 is observing ant he has not   *)
1749(* paid. this is fine because the ring is symmetric and the numbering is     *)
1750(* arbitrary.                                                                *)
1751(* ************************************************************************* *)
1752
1753val insider_high_states_set_def = Define   `(insider_high_states_set (0:num) = {}) /\    (insider_high_states_set (SUC n) = (\s:string. s = (STRCAT "pays" (toString (SUC n))))
1754                                                        INSERT (insider_high_states_set n))`;val insider_high_states_def = Define   `insider_high_states nsapays (SUC(SUC n)) = if nsapays then                                         {(\s: string. s = STRCAT "pays" (toString (SUC(SUC n))))}                                         else                                          insider_high_states_set (SUC n)`;val insider_low_states_def = Define   `insider_low_states = {(\s:string. s = (STRCAT "coin" (toString 0))); (\s:string. F)}`;
1755
1756val insider_random_states_def = Define   `(insider_random_states (0:num) = {(\s:string. F)}) /\    (insider_random_states (SUC n) = (IMAGE (\s:bool state.                                            (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then                                                                T                                                         else s x))                                      (insider_random_states n)) UNION                                   (IMAGE (\s:bool state.                                           (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then                                                                F                                                         else s x))                                      (insider_random_states n)))`;
1757
1758val insider_set_announcements_def = Define   `(insider_set_announcements (high: bool state) (low:bool state) (random:bool state) (n:num) (0:num) (s:string) =           if (s = (STRCAT "announces" (toString 0))) then                 ((high (STRCAT "pays" (toString 0)))) xor                       ((low (STRCAT "coin" (toString 0))) xor (random (STRCAT "coin" (toString n))))          else                    low s) /\
1759    (insider_set_announcements (high: bool state) (low:bool state) (random:bool state) (n:num) (SUC(0:num)) (s:string) =                if (s = (STRCAT "announces" (toString (SUC 0)))) then                   ((high (STRCAT "pays" (toString (SUC 0))))) xor                 ((random (STRCAT "coin" (toString (SUC 0)))) xor (low (STRCAT "coin" (toString 0))))            else                    (insider_set_announcements high low random n 0) s) /\    (insider_set_announcements high low random n (SUC (SUC i)) s =         if (s = (STRCAT "announces" (toString (SUC (SUC i))))) then                     ((high (STRCAT "pays" (toString (SUC (SUC i)))))) xor                   ((random (STRCAT "coin" (toString (SUC (SUC i))))) xor (random (STRCAT "coin" (toString (SUC i)))))             else                    (insider_set_announcements high low random n (SUC i)) s)`;
1760
1761val insider_set_announcements_alt = store_thm  ("insider_set_announcements_alt",   ``!high low random n i.      (insider_set_announcements high low random n 0 =         (\s. if (s = (STRCAT "announces" (toString 0))) then                   ((high (STRCAT "pays" (toString 0)))) xor                       ((low (STRCAT "coin" (toString 0))) xor (random (STRCAT "coin" (toString n))))          else                    low s)) /\
1762        (insider_set_announcements high low random n (SUC 0) =   (\s. if (s = (STRCAT "announces" (toString (SUC 0)))) then                     ((high (STRCAT "pays" (toString (SUC 0))))) xor                 ((random (STRCAT "coin" (toString (SUC 0)))) xor (low (STRCAT "coin" (toString 0))))            else                    (insider_set_announcements high low random n 0) s)) /\  (insider_set_announcements high low random n (SUC(SUC i)) =      (\s. if (s = (STRCAT "announces" (toString (SUC(SUC i))))) then                        ((high (STRCAT "pays" (toString (SUC(SUC i)))))) xor                    ((random (STRCAT "coin" (toString (SUC(SUC i))))) xor (random (STRCAT "coin" (toString (SUC i)))))              else                    (insider_set_announcements high low random n (SUC i)) s))``,
1763  RW_TAC std_ss [FUN_EQ_THM, insider_set_announcements_def]);
1764val insider_dcprog_def = Define   `insider_dcprog (SUC(SUC(SUC n))) = (\((high:bool state, low:bool state), random:bool state). compute_result (insider_set_announcements high low random (SUC(SUC n)) (SUC(SUC n)))                   (SUC(SUC n)))`;
1765
1766val insider_dc_prog_space_def = Define   `insider_dc_prog_space (SUC(SUC n)) nsapays =  unif_prog_space         (insider_high_states nsapays (SUC(SUC n)))              insider_low_states              (insider_random_states (SUC n))`;
1767
1768val insider_dc_prog_space_F_set_thm = store_thm
1769  ("insider_dc_prog_space_F_set_thm",
1770   ``!n. insider_dc_prog_space (SUC(SUC n)) F =
1771         unif_prog_space (insider_high_states_set (SUC n))
1772                         insider_low_states
1773                         (insider_random_states (SUC n))``,
1774   RW_TAC std_ss [insider_dc_prog_space_def, insider_high_states_def]);
1775(* ------------------------------------------------------------------------- *)
1776
1777val hl_dups_lemma = prove
1778   (``!(s1:string)(s2:string). ((\x. x = s1) = (\x. x = s2)) = (s1=s2)``,
1779    METIS_TAC []);
1780
1781val hl_dups_lemma2 = prove
1782   (``!(s:string). ~((\x. x = s) = (\x. F))``,
1783    METIS_TAC []);
1784
1785
1786val fun_eq_lem6 = METIS_PROVE []
1787        ``!x a b P Q. (x = a) \/ ~ (x = b) /\ P = (a = b) /\ ((x = a) \/ P) \/ ~(x = b) /\ ((x = a) \/ P)``;
1788
1789fun insider_input_unroll (remove_dups_conv:Abbrev.term->Abbrev.thm) =
1790        (REPEATC (SIMP_CONV bool_ss [insider_high_states_set_def, insider_low_states_def, insider_random_states_def,
1791                                     CARD_DEF, FINITE_INSERT, FINITE_EMPTY, FINITE_SING,
1792                                     IMAGE_UNION, IMAGE_IMAGE, combinTheory.o_DEF, IMAGE_INSERT, IMAGE_EMPTY,
1793                                     INSERT_UNION, UNION_EMPTY, IN_UNION]
1794         THENC (FIND_CONV ``x IN y`` (IN_CONV remove_dups_conv)
1795                THENC SIMP_CONV bool_ss [])
1796         THENC SIMP_CONV arith_ss []));
1797
1798val dc_dup_conv = SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
1799                                              THENC EVAL
1800                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
1801                                              THENC EVAL
1802                                              THENC (REPEATC (T_F_UNCHANGED_CONV
1803                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
1804                                                         THENC EVAL)));
1805
1806val insider_hl_dup_conv = SIMP_CONV arith_ss [hl_dups_lemma, hl_dups_lemma2, STRCAT_toString_inj];
1807
1808val insider_dc3_leakage_result = store_thm
1809  ("insider_dc3_leakage_result",
1810   ``leakage (insider_dc_prog_space (SUC (SUC (SUC 0))) F) (insider_dcprog (SUC(SUC(SUC 0)))) = 0``,
1811   CONV_TAC (SIMP_CONV bool_ss [insider_dc_prog_space_F_set_thm] THENC
1812             RATOR_CONV (RAND_CONV (
1813             LEAKAGE_COMPUTE_CONV (``insider_high_states_set (SUC(SUC 0))``,``insider_low_states``,``insider_random_states (SUC(SUC 0))``)
1814[insider_high_states_set_def, insider_low_states_def, insider_random_states_def]
1815[insider_dcprog_def, compute_result_alt, XOR_announces_def, insider_set_announcements_alt, FST, SND, xor_def]
1816(insider_input_unroll insider_hl_dup_conv) (insider_input_unroll insider_hl_dup_conv) (insider_input_unroll dc_dup_conv)
1817insider_hl_dup_conv insider_hl_dup_conv dc_dup_conv dc_dup_conv))
1818             THENC SIMP_CONV real_ss [REAL_ADD_ASSOC, GSYM REAL_NEG_ADD, GSYM REAL_ADD_RDISTRIB, REAL_MUL_ASSOC])
1819>> RW_TAC real_ss [lg_1, lg_inv, GSYM REAL_INV_1OVER]
1820>> `lg 4 = 2` by (`4 = 2 pow 2` by RW_TAC real_ss [pow] >> POP_ORW >> RW_TAC std_ss [lg_pow])
1821>> RW_TAC real_ss []);
1822
1823(* ************************************************************************* *)
1824(* Case Study: The Dining Cryptographers - Biased coins                      *)
1825(* ************************************************************************* *)
1826
1827(* ************************************************************************* *)
1828(* Tampering w/ a visible coin doesn't create a leak                         *)
1829(* ************************************************************************* *)
1830
1831val biased_high_states_def = Define
1832   `biased_high_states n = insider_high_states_set n`;
1833
1834val biased_low_states_def = Define
1835   `biased_low_states = {(\s:string. s = (STRCAT "coin" (toString 0))); (\s:string. F)}`;
1836
1837val biased_random_states_def = Define
1838   `(biased_random_states (0:num) = {(\s:string. F)}) /\
1839    (biased_random_states (SUC n) = (IMAGE (\s:bool state.
1840                                              (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then
1841                                                                T
1842                                                          else s x))
1843                                          (biased_random_states n))
1844                                        UNION
1845                                     (IMAGE (\s:bool state.
1846                                              (\x:string. if x = (STRCAT "coin" (toString (SUC n))) then
1847                                                                F
1848                                                          else s x))
1849                                          (biased_random_states n)))`;
1850
1851val biased_dist_def = Define
1852   `biased_dist high low random =
1853        (\s. if L s = (\s:string. s = (STRCAT "coin" (toString 0))) then
1854                (1 / 2) * (unif_prog_dist high low random s)
1855             else
1856                (3 / 2) * (unif_prog_dist high low random s))`;
1857
1858val biased_dc_prog_space_def = Define   `biased_dc_prog_space (SUC(SUC n)) =
1859        (((biased_high_states (SUC(SUC n))) CROSS biased_low_states) CROSS (biased_random_states (SUC (SUC n))),
1860         POW (((biased_high_states (SUC(SUC n))) CROSS biased_low_states) CROSS (biased_random_states (SUC (SUC n)))),
1861         (\s. SIGMA (biased_dist (biased_high_states (SUC(SUC n))) biased_low_states (biased_random_states (SUC (SUC n)))) s))`;
1862val prob_space_biased_dc_prog_space3 = store_thm
1863  ("prob_space_biased_dc_prog_space3",
1864   ``prob_space (biased_dc_prog_space (SUC(SUC 0)))``,
1865   (RW_TAC bool_ss [prob_space_def]
1866    >> `FINITE ((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1867                biased_random_states (SUC (SUC 0))))`
1868        by RW_TAC set_ss [biased_high_states_def, biased_low_states_def,
1869                             biased_random_states_def, insider_high_states_set_def])
1870   >- (MATCH_MP_TAC finite_additivity_sufficient_for_finite_spaces2
1871       >> RW_TAC bool_ss [biased_dc_prog_space_def, m_space_def, measurable_sets_def, POW_SIGMA_ALGEBRA,
1872                          PSPACE, measure_def, positive_def, REAL_SUM_IMAGE_THM, IN_POW,
1873                          additive_def, IN_FUNSET, IN_UNIV]
1874       >- (MATCH_MP_TAC REAL_SUM_IMAGE_POS
1875           >> CONJ_TAC
1876           >- METIS_TAC [SUBSET_FINITE]
1877           >> RW_TAC real_ss [biased_dist_def, unif_prog_dist_def, GSYM REAL_INV_1OVER, REAL_LE_INV_EQ, REAL_LE_MUL])
1878       >> METIS_TAC [SUBSET_FINITE, REAL_SUM_IMAGE_DISJOINT_UNION])
1879   >> RW_TAC bool_ss [biased_dc_prog_space_def, m_space_def, measurable_sets_def,
1880                      PSPACE, measure_def, biased_dist_def]
1881   >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1882        biased_random_states (SUC (SUC 0))) =
1883        (((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1884         biased_random_states (SUC (SUC 0))))INTER
1885         {s| L s = (\s. s = STRCAT "coin" (toString 0))}) UNION
1886        (((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1887        biased_random_states (SUC (SUC 0))))INTER{s| ~(L s = (\s. s = STRCAT "coin" (toString 0)))})`
1888        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, IN_INTER, GSPECIFICATION] >> DECIDE_TAC)
1889   >> POP_ORW
1890   >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1891        biased_random_states (SUC (SUC 0)) INTER
1892        {s | L s = (\s. s = STRCAT "coin" (toString 0))}) /\
1893        FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1894        biased_random_states (SUC (SUC 0)) INTER
1895        {s | ~(L s = (\s. s = STRCAT "coin" (toString 0)))})`
1896        by RW_TAC std_ss [INTER_FINITE]
1897   >> `DISJOINT (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1898        biased_random_states (SUC (SUC 0)) INTER
1899        {s | L s = (\s. s = STRCAT "coin" (toString 0))})
1900        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1901        biased_random_states (SUC (SUC 0)) INTER
1902        {s | ~(L s = (\s. s = STRCAT "coin" (toString 0)))})`
1903        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
1904            >> DECIDE_TAC)
1905   >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION]
1906   >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF, unif_prog_dist_def]
1907   >> `(\s.
1908     (if
1909        s IN
1910        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1911        biased_random_states (SUC (SUC 0)) INTER
1912        {s | L s = (\s. s = STRCAT "coin" (toString 0))}
1913      then
1914        (if L s = (\s. s = STRCAT "coin" (toString 0)) then
1915           1 / 2 *
1916           (if
1917              s IN
1918              biased_high_states (SUC (SUC 0)) CROSS
1919              biased_low_states CROSS biased_random_states (SUC (SUC 0))
1920            then
1921              1 /
1922              &
1923                (CARD
1924                   (biased_high_states (SUC (SUC 0)) CROSS
1925                    biased_low_states CROSS
1926                    biased_random_states (SUC (SUC 0))))
1927            else
1928              0)
1929         else
1930           3 / 2 *
1931           (if
1932              s IN
1933              biased_high_states (SUC (SUC 0)) CROSS
1934              biased_low_states CROSS biased_random_states (SUC (SUC 0))
1935            then
1936              1 /
1937              &
1938                (CARD
1939                   (biased_high_states (SUC (SUC 0)) CROSS
1940                    biased_low_states CROSS
1941                    biased_random_states (SUC (SUC 0))))
1942            else
1943              0))
1944      else
1945        0)) =
1946        (\s. (if s IN
1947        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1948        biased_random_states (SUC (SUC 0)) INTER
1949        {s | L s = (\s. s = STRCAT "coin" (toString 0))}
1950      then
1951        (\s. 1 / 2 * (1 /
1952              &
1953                (CARD
1954                   (biased_high_states (SUC (SUC 0)) CROSS
1955                    biased_low_states CROSS
1956                    biased_random_states (SUC (SUC 0)))))) s
1957      else 0))`
1958        by (ONCE_REWRITE_TAC [FUN_EQ_THM]
1959            >> RW_TAC bool_ss [IN_INTER, GSPECIFICATION])
1960   >> POP_ORW
1961   >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
1962   >> RW_TAC bool_ss [REAL_SUM_IMAGE_FINITE_CONST3]
1963   >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
1964   >> `(\s.
1965     (if
1966        s IN
1967        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
1968        biased_random_states (SUC (SUC 0)) INTER
1969        {s | ~(L s = (\s. s = STRCAT "coin" (toString 0)))}
1970      then
1971        (if L s = (\s. s = STRCAT "coin" (toString 0)) then
1972           1 / 2 *
1973           (if
1974              s IN
1975              biased_high_states (SUC (SUC 0)) CROSS
1976              biased_low_states CROSS biased_random_states (SUC (SUC 0))
1977            then
1978              1 /
1979              &
1980                (CARD
1981                   (biased_high_states (SUC (SUC 0)) CROSS
1982                    biased_low_states CROSS
1983                    biased_random_states (SUC (SUC 0))))
1984            else
1985              0)
1986         else
1987           3/2 *
1988           (if
1989              s IN
1990              biased_high_states (SUC (SUC 0)) CROSS
1991              biased_low_states CROSS biased_random_states (SUC (SUC 0))
1992            then
1993              1 /
1994              &
1995                (CARD
1996                   (biased_high_states (SUC (SUC 0)) CROSS
1997                    biased_low_states CROSS
1998                    biased_random_states (SUC (SUC 0))))
1999            else
2000              0))
2001      else
2002        0)) =
2003        (\s. (if s IN
2004        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2005        biased_random_states (SUC (SUC 0)) INTER
2006        {s | ~(L s = (\s. s = STRCAT "coin" (toString 0)))}
2007      then
2008        (\s. 3/2 * (1 /
2009              &
2010                (CARD
2011                   (biased_high_states (SUC (SUC 0)) CROSS
2012                    biased_low_states CROSS
2013                    biased_random_states (SUC (SUC 0)))))) s
2014      else 0))`
2015        by (ONCE_REWRITE_TAC [FUN_EQ_THM]
2016            >> RW_TAC bool_ss [IN_INTER, GSPECIFICATION])
2017   >> POP_ORW
2018   >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
2019   >> RW_TAC bool_ss [REAL_SUM_IMAGE_FINITE_CONST3]
2020   >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2021      biased_random_states (SUC (SUC 0)) INTER
2022      {s | L s = (\s. s = STRCAT "coin" (toString 0))}) =
2023       (biased_high_states (SUC (SUC 0)) CROSS {(\s. s = STRCAT "coin" (toString 0))} CROSS
2024      biased_random_states (SUC (SUC 0)))`
2025        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [IN_INTER, L_def, GSPECIFICATION, IN_CROSS, IN_SING, biased_low_states_def]
2026            >> DECIDE_TAC)
2027   >> POP_ORW
2028   >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2029      biased_random_states (SUC (SUC 0)) INTER
2030      {s | ~(L s = (\s. s = STRCAT "coin" (toString 0)))}) =
2031       (biased_high_states (SUC (SUC 0)) CROSS {(\s:string. F)} CROSS
2032      biased_random_states (SUC (SUC 0)))`
2033        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [IN_INTER, L_def, GSPECIFICATION, IN_CROSS, IN_SING, biased_low_states_def]
2034            >> METIS_TAC [])
2035   >> POP_ORW
2036  >> RW_TAC bool_ss [(COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS
2037                                {(\s. s = STRCAT "coin" (toString 0))} CROSS
2038                                biased_random_states (SUC (SUC 0)))``
2039                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
2040                                biased_random_states_def, insider_high_states_set_def] THENC
2041                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2042                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2043                                                     THENC (TRY_CONV insider_hl_dup_conv)
2044                                                     THENC (TRY_CONV insider_hl_dup_conv))))
2045                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2046         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2047                                                     THENC (TRY_CONV insider_hl_dup_conv)
2048                                                     THENC (TRY_CONV insider_hl_dup_conv)
2049                                                     THENC (TRY_CONV dc_dup_conv)))))
2050                dc_dup_conv),
2051                (COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2052                                biased_random_states (SUC (SUC 0)))``
2053                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
2054                                biased_random_states_def, insider_high_states_set_def] THENC
2055                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2056                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2057                                                     THENC (TRY_CONV insider_hl_dup_conv)
2058                                                     THENC (TRY_CONV insider_hl_dup_conv))))
2059                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2060         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2061                                                     THENC (TRY_CONV insider_hl_dup_conv)
2062                                                     THENC (TRY_CONV insider_hl_dup_conv)
2063                                                     THENC (TRY_CONV dc_dup_conv)))))
2064                dc_dup_conv),
2065        (COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS {(\s. F)} CROSS
2066                        biased_random_states (SUC (SUC 0)))``
2067                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
2068                                biased_random_states_def, insider_high_states_set_def] THENC
2069                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2070                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2071                                                     THENC (TRY_CONV insider_hl_dup_conv)
2072                                                     THENC (TRY_CONV insider_hl_dup_conv))))
2073                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2074         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2075                                                     THENC (TRY_CONV insider_hl_dup_conv)
2076                                                     THENC (TRY_CONV insider_hl_dup_conv)
2077                                                     THENC (TRY_CONV dc_dup_conv)))))
2078                dc_dup_conv)]
2079   >> RW_TAC real_ss []);
2080
2081(* ------------------------------------------------------------------------- *)
2082
2083val biased_dc3_leakage_result = store_thm
2084  ("biased_dc3_leakage_result",
2085   ``(leakage (biased_dc_prog_space (SUC (SUC 0))) (insider_dcprog (SUC(SUC(SUC 0)))) = 0)``,
2086   RW_TAC bool_ss [leakage_def]
2087   >> `FINITE ((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2088                biased_random_states (SUC (SUC 0))))`
2089        by RW_TAC set_ss [biased_high_states_def, biased_low_states_def,
2090                             biased_random_states_def, insider_high_states_set_def]
2091   >> `FINITE (biased_high_states (SUC (SUC 0)))`
2092        by RW_TAC set_ss [biased_high_states_def, insider_high_states_set_def]
2093   >> `FINITE (biased_random_states (SUC (SUC 0)))`
2094        by RW_TAC set_ss [biased_random_states_def]
2095   >> `FINITE (biased_low_states)`
2096        by RW_TAC set_ss [biased_low_states_def]
2097   >> (MP_TAC o
2098       Q.SPECL [`2`,`(biased_dc_prog_space (SUC (SUC 0)))`, `(insider_dcprog (SUC (SUC (SUC 0))))`, `H`,`L`] o
2099       INST_TYPE [alpha |-> ``:(bool, bool, bool) prog_state``, beta |-> ``:bool state``, gamma |-> ``:bool state``,
2100                  delta |-> ``:bool state``])
2101        finite_conditional_mutual_information_reduce
2102   >> (MP_TAC o INST_TYPE [beta |-> ``:bool state``] o Q.ISPEC `(biased_dc_prog_space (SUC (SUC 0)))`)
2103        MEASURABLE_POW_TO_POW_IMAGE
2104   >> RW_TAC bool_ss [random_variable_def, prob_space_biased_dc_prog_space3, biased_dc_prog_space_def, PSPACE, EVENTS,
2105                      prob_space_def, measurable_sets_def, m_space_def,
2106                      REWRITE_RULE [prob_space_def] prob_space_biased_dc_prog_space3]
2107   >> POP_ASSUM (K ALL_TAC)
2108   >> RW_TAC bool_ss [biased_dist_def, unif_prog_dist_def,
2109        (COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
2110                                biased_random_states (SUC (SUC 0)))``
2111                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
2112                                biased_random_states_def, insider_high_states_set_def] THENC
2113                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2114                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2115                                                     THENC (TRY_CONV insider_hl_dup_conv)
2116                                                     THENC (TRY_CONV insider_hl_dup_conv))))
2117                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2118         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2119                                                     THENC (TRY_CONV insider_hl_dup_conv)
2120                                                     THENC (TRY_CONV insider_hl_dup_conv)
2121                                                     THENC (TRY_CONV dc_dup_conv)))))
2122                dc_dup_conv)]
2123   >> Q.ABBREV_TAC `foo = (SUC (SUC 0))`
2124   >> RW_TAC std_ss [joint_distribution_def, L_def, H_def, PAIR, ETA_THM, FST, SND, PROB, PSPACE, distribution_def]
2125   >> `IMAGE L
2126        (biased_high_states foo CROSS biased_low_states CROSS
2127         biased_random_states foo) =
2128        biased_low_states`
2129        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, L_def]
2130            >> EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []
2131            >> Q.UNABBREV_TAC `foo`
2132            >> Q.EXISTS_TAC `(((\s:string. s = STRCAT "pays" (toString (SUC (SUC 0)))),x),(\s:string. F))`
2133            >> RW_TAC bool_ss [SND, FST, biased_high_states_def, insider_high_states_set_def, IN_INSERT]
2134            >> Suff `!n. (\s:string. F) IN biased_random_states n` >- RW_TAC std_ss []
2135            >> REPEAT (POP_ASSUM (K ALL_TAC))
2136            >> Induct >> RW_TAC bool_ss [biased_random_states_def, IN_SING, IN_IMAGE, IN_UNION]
2137            >> DISJ2_TAC >> Q.EXISTS_TAC `(\s:string.F)` >> RW_TAC std_ss [])
2138   >> POP_ORW
2139   >> `(IMAGE (insider_dcprog (SUC foo))
2140        (biased_high_states foo CROSS biased_low_states CROSS
2141         biased_random_states foo) CROSS
2142      IMAGE FST
2143        (biased_high_states foo CROSS biased_low_states CROSS
2144         biased_random_states foo)) =
2145        (IMAGE (insider_dcprog (SUC foo))
2146        (biased_high_states foo CROSS biased_low_states CROSS
2147         biased_random_states foo) CROSS
2148      (biased_high_states foo CROSS biased_low_states))`
2149        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, FST]
2150            >> EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []
2151            >> Q.EXISTS_TAC `(SND x, SND x')`
2152            >> RW_TAC bool_ss [SND, FST])
2153   >> POP_ORW
2154   >> `!x z.
2155        SIGMA
2156          (\s.
2157             (if SND (FST s) = (\s. s = STRCAT "coin" (toString 0)) then
2158                1 / 2 *
2159                (if
2160                   s IN
2161                   biased_high_states foo CROSS biased_low_states CROSS
2162                   biased_random_states foo
2163                 then
2164                   1 / 16
2165                 else
2166                   0)
2167              else
2168                3 / 2 *
2169                (if
2170                   s IN
2171                   biased_high_states foo CROSS biased_low_states CROSS
2172                   biased_random_states foo
2173                 then
2174                   1 / 16
2175                 else
2176                   0)))
2177          (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2178           PREIMAGE L {z} INTER
2179           (biased_high_states foo CROSS biased_low_states CROSS
2180            biased_random_states foo)) =
2181        (if (z = (\s. s = STRCAT "coin" (toString 0))) then
2182                1/32
2183         else (if (z = (\s:string. F)) then
2184                3/32
2185         else 0)) *
2186        &(CARD {(h,r)| h IN biased_high_states foo /\ r IN biased_random_states foo /\
2187                            (insider_dcprog (SUC foo) ((h,z),r) = x)})`
2188        by (RW_TAC real_ss []
2189            >| [`(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2190                 PREIMAGE L {(\s. s = STRCAT "coin" (toString 0))} INTER
2191                (biased_high_states foo CROSS biased_low_states CROSS
2192                 biased_random_states foo)) =
2193                 IMAGE (\(h,r). ((h, (\s. s = STRCAT "coin" (toString 0))),r))
2194                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2195                                (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)}`
2196                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2197                                    IN_CROSS, L_def]
2198                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2199                     >- (Q.PAT_X_ASSUM `P = (\s:string. s = STRCAT "coin" (toString 0))` (ASSUME_TAC o GSYM)
2200                         >> Q.EXISTS_TAC `(FST(FST x'),SND x')`
2201                         >> RW_TAC std_ss [PAIR]
2202                         >> Q.EXISTS_TAC `(FST(FST x'),SND x')`
2203                         >> RW_TAC std_ss [PAIR_EQ])
2204                     >> POP_ASSUM MP_TAC
2205                     >> `x''' = (FST x''',SND x''')` by RW_TAC std_ss [PAIR]
2206                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
2207                     >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2208                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND]
2209                     >> RW_TAC set_ss [biased_low_states_def])
2210                >> POP_ORW
2211                >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2212                                        (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)}`
2213                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
2214                       >> RW_TAC bool_ss [FINITE_CROSS]
2215                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
2216                       >> POP_ASSUM MP_TAC
2217                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2218                       >> POP_ORW >> RW_TAC std_ss [])
2219                >> `INJ (\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r))
2220                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2221                                 (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2222                        (IMAGE (\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r))
2223                                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2224                                        (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)})`
2225                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE] >- METIS_TAC []
2226                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2227                            >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2228                            >> POP_ORW >> RW_TAC std_ss []
2229                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2230                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
2231                            >> POP_ORW >> RW_TAC std_ss []
2232                            >> METIS_TAC [PAIR, PAIR_EQ])
2233                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF] >> POP_ASSUM (K ALL_TAC)
2234                >> `!(x' :bool state # bool state).
2235                         SND (FST ((\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r)) x')) =
2236                         (\s. s = STRCAT "coin" (toString 0))`
2237                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2238                            >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
2239                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2240                >> `!(x' :bool state # bool state).
2241                         ((\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r)) x' IN
2242                        biased_high_states foo CROSS biased_low_states CROSS
2243                        biased_random_states foo) = (FST x' IN (biased_high_states foo) /\
2244                                                     SND x' IN (biased_random_states foo))`
2245                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2246                            >> POP_ORW >> RW_TAC std_ss []
2247                            >> RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND])
2248                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2249                >> `(\(x' :bool state # bool state). (if x' IN
2250                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2251                                (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2252                        then 1 / 2 * (if FST x' IN biased_high_states foo /\ SND x' IN biased_random_states foo
2253                        then 1 / 16 else 0) else 0)) =
2254                    (\(x' :bool state # bool state). (if x' IN
2255                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2256                                (insider_dcprog (SUC foo) ((h,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2257                        then (\x. 1 / 32) x' else 0))`
2258                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [] >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2259                            >> RW_TAC std_ss [GSPECIFICATION]
2260                            >> POP_ASSUM MP_TAC >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2261                            >> POP_ORW >> RW_TAC std_ss [])
2262                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM],
2263                `(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2264                 PREIMAGE L {(\s:string. F)} INTER
2265                (biased_high_states foo CROSS biased_low_states CROSS
2266                 biased_random_states foo)) =
2267                 IMAGE (\(h,r). ((h, (\s:string. F)),r))
2268                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2269                                (insider_dcprog (SUC foo) ((h,(\s:string. F)),r) = x)}`
2270                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2271                                    IN_CROSS, L_def]
2272                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2273                     >- (Q.PAT_X_ASSUM `P = (\s:string. F)` (ASSUME_TAC o GSYM)
2274                         >> Q.EXISTS_TAC `(FST(FST x'),SND x')`
2275                         >> RW_TAC std_ss [PAIR]
2276                         >> Q.EXISTS_TAC `(FST(FST x'),SND x')`
2277                         >> RW_TAC std_ss [PAIR_EQ])
2278                     >> POP_ASSUM MP_TAC
2279                     >> `x''' = (FST x''',SND x''')` by RW_TAC std_ss [PAIR]
2280                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
2281                     >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2282                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND]
2283                     >> RW_TAC set_ss [biased_low_states_def])
2284                >> POP_ORW
2285                >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2286                                        (insider_dcprog (SUC foo) ((h,(\s:string. F)),r) = x)}`
2287                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
2288                       >> RW_TAC bool_ss [FINITE_CROSS]
2289                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
2290                       >> POP_ASSUM MP_TAC
2291                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2292                       >> POP_ORW >> RW_TAC std_ss [])
2293                >> `INJ (\(h,r). ((h,(\s:string. F)),r))
2294                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2295                                 (insider_dcprog (SUC foo) ((h,(\s:string. F)),r) = x)}
2296                        (IMAGE (\(h,r). ((h,(\s:string. F)),r))
2297                                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2298                                        (insider_dcprog (SUC foo) ((h,(\s:string. F)),r) = x)})`
2299                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE] >- METIS_TAC []
2300                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2301                            >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2302                            >> POP_ORW >> RW_TAC std_ss []
2303                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2304                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
2305                            >> POP_ORW >> RW_TAC std_ss []
2306                            >> METIS_TAC [PAIR, PAIR_EQ])
2307                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF] >> POP_ASSUM (K ALL_TAC)
2308                >> `!(x' :bool state # bool state).
2309                         SND (FST ((\(h,r). ((h,(\s:string.F)),r)) x')) =
2310                         (\s:string. F)`
2311                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2312                            >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
2313                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2314                >> `!(x' :bool state # bool state).
2315                         ((\(h,r). ((h,(\s:string.F)),r)) x' IN
2316                        biased_high_states foo CROSS biased_low_states CROSS
2317                        biased_random_states foo) = (FST x' IN (biased_high_states foo) /\
2318                                                     SND x' IN (biased_random_states foo))`
2319                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2320                            >> POP_ORW >> RW_TAC std_ss []
2321                            >> RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND])
2322                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2323                >> `(\(x' :bool state # bool state). (if x' IN
2324                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2325                                (insider_dcprog (SUC foo) ((h,(\s:string.F)),r) = x)}
2326                        then 3 / 2 * (if FST x' IN biased_high_states foo /\ SND x' IN biased_random_states foo
2327                        then 1 / 16 else 0) else 0)) =
2328                    (\(x' :bool state # bool state). (if x' IN
2329                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
2330                                (insider_dcprog (SUC foo) ((h,(\s:string.F)),r) = x)}
2331                        then (\x. 3 / 32) x' else 0))`
2332                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [] >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2333                            >> RW_TAC std_ss [GSPECIFICATION]
2334                            >> POP_ASSUM MP_TAC >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
2335                            >> POP_ORW >> RW_TAC std_ss [])
2336                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM],
2337                RW_TAC bool_ss [GSYM INTER_ASSOC]
2338                >> Suff `PREIMAGE L {z} INTER (biased_high_states foo CROSS biased_low_states CROSS
2339                                                biased_random_states foo) = {}`
2340                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
2341                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, L_def, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING, biased_low_states_def]
2342                >> METIS_TAC []])
2343   >> `!x z. (PREIMAGE (\x. (insider_dcprog (SUC foo) x,SND (FST x)))
2344           {(x,z)} INTER
2345         (biased_high_states foo CROSS biased_low_states CROSS
2346          biased_random_states foo)) =
2347        (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2348           PREIMAGE L {z} INTER
2349           (biased_high_states foo CROSS biased_low_states CROSS
2350            biased_random_states foo))`
2351        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, IN_CROSS, IN_PREIMAGE, IN_SING, L_def])
2352   >> POP_ORW
2353   >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2354   >> `!x z.
2355        SIGMA
2356             (\s.
2357                (if
2358                   SND (FST s) = (\s. s = STRCAT "coin" (toString 0))
2359                 then
2360                   1 / 2 *
2361                   (if
2362                      s IN
2363                      biased_high_states foo CROSS
2364                      biased_low_states CROSS biased_random_states foo
2365                    then
2366                      1 / 16
2367                    else
2368                      0)
2369                 else
2370                   3 / 2 *
2371                   (if
2372                      s IN
2373                      biased_high_states foo CROSS
2374                      biased_low_states CROSS biased_random_states foo
2375                    then
2376                      1 / 16
2377                    else
2378                      0)))
2379             (PREIMAGE L {z} INTER
2380              (biased_high_states foo CROSS biased_low_states CROSS
2381               biased_random_states foo)) =
2382        (if (z = (\s. s = STRCAT "coin" (toString 0))) then
2383                1/32
2384         else (if (z = (\s:string. F)) then
2385                3/32
2386         else 0)) *
2387        &(CARD (biased_high_states foo) * CARD (biased_random_states foo))`
2388        by (RW_TAC real_ss []
2389            >| [`(PREIMAGE L {(\s. s = STRCAT "coin" (toString 0))} INTER
2390                (biased_high_states foo CROSS biased_low_states CROSS
2391                 biased_random_states foo)) =
2392                 IMAGE (\(h,r). ((h, (\s. s = STRCAT "coin" (toString 0))),r))
2393                        (biased_high_states foo CROSS biased_random_states foo)`
2394                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2395                                    IN_CROSS, L_def]
2396                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2397                     >- (Q.PAT_X_ASSUM `P = (\s:string. s = STRCAT "coin" (toString 0))` (ASSUME_TAC o GSYM)
2398                         >> Q.EXISTS_TAC `(FST(FST x),SND x)`
2399                         >> RW_TAC std_ss [PAIR]
2400                         >> Q.EXISTS_TAC `(FST(FST x),SND x)`
2401                         >> RW_TAC std_ss [PAIR_EQ])
2402                     >>  `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2403                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND]
2404                     >> RW_TAC set_ss [biased_low_states_def])
2405                >> POP_ORW
2406                >> `INJ (\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r))
2407                        (biased_high_states foo CROSS biased_random_states foo)
2408                        (IMAGE (\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r))
2409                                (biased_high_states foo CROSS biased_random_states foo))`
2410                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE] >- METIS_TAC []
2411                            >> POP_ASSUM MP_TAC
2412                            >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
2413                            >> POP_ORW >> RW_TAC std_ss []
2414                            >> POP_ASSUM MP_TAC
2415                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
2416                            >> POP_ORW >> RW_TAC std_ss []
2417                            >> METIS_TAC [PAIR, PAIR_EQ])
2418                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF, FINITE_CROSS] >> POP_ASSUM (K ALL_TAC)
2419                >> `!(x' :bool state # bool state).
2420                         SND (FST ((\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r)) x')) =
2421                         (\s. s = STRCAT "coin" (toString 0))`
2422                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2423                            >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
2424                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2425                >> `!(x' :bool state # bool state).
2426                         ((\(h,r). ((h,(\s. s = STRCAT "coin" (toString 0))),r)) x' IN
2427                        biased_high_states foo CROSS biased_low_states CROSS
2428                        biased_random_states foo) = (FST x' IN (biased_high_states foo) /\
2429                                                     SND x' IN (biased_random_states foo))`
2430                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2431                            >> POP_ORW >> RW_TAC std_ss []
2432                            >> RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND])
2433                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2434                >> `(\(x' :bool state # bool state). (if x' IN
2435                        biased_high_states foo CROSS biased_random_states foo
2436                        then 1 / 2 * (if FST x' IN biased_high_states foo /\ SND x' IN biased_random_states foo
2437                        then 1 / 16 else 0) else 0)) =
2438                    (\(x' :bool state # bool state). (if x' IN
2439                        biased_high_states foo CROSS biased_random_states foo
2440                        then (\x. 1 / 32) x' else 0))`
2441                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [IN_CROSS])
2442                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3,
2443                                              REAL_MUL_COMM, FINITE_CROSS, CARD_CROSS],
2444                `(PREIMAGE L {(\s:string. F)} INTER
2445                (biased_high_states foo CROSS biased_low_states CROSS
2446                 biased_random_states foo)) =
2447                 IMAGE (\(h,r). ((h, (\s:string. F)),r))
2448                        (biased_high_states foo CROSS biased_random_states foo)`
2449                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2450                                    IN_CROSS, L_def]
2451                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2452                     >- (Q.PAT_X_ASSUM `P = (\s:string. F)` (ASSUME_TAC o GSYM)
2453                         >> Q.EXISTS_TAC `(FST(FST x),SND x)`
2454                         >> RW_TAC std_ss [PAIR]
2455                         >> Q.EXISTS_TAC `(FST(FST x),SND x)`
2456                         >> RW_TAC std_ss [PAIR_EQ])
2457                     >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2458                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND]
2459                     >> RW_TAC set_ss [biased_low_states_def])
2460                >> POP_ORW
2461                >> `INJ (\(h,r). ((h,(\s:string. F)),r))
2462                        (biased_high_states foo CROSS biased_random_states foo)
2463                        (IMAGE (\(h,r). ((h,(\s:string. F)),r))
2464                                (biased_high_states foo CROSS biased_random_states foo))`
2465                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE] >- METIS_TAC []
2466                            >> POP_ASSUM MP_TAC
2467                            >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
2468                            >> POP_ORW >> RW_TAC std_ss []
2469                            >> POP_ASSUM MP_TAC
2470                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
2471                            >> POP_ORW >> RW_TAC std_ss []
2472                            >> METIS_TAC [PAIR, PAIR_EQ])
2473                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF, FINITE_CROSS] >> POP_ASSUM (K ALL_TAC)
2474                >> `!(x' :bool state # bool state).
2475                         SND (FST ((\(h,r). ((h,(\s:string.F)),r)) x')) =
2476                         (\s:string. F)`
2477                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2478                            >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
2479                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2480                >> `!(x' :bool state # bool state).
2481                         ((\(h,r). ((h,(\s:string.F)),r)) x' IN
2482                        biased_high_states foo CROSS biased_low_states CROSS
2483                        biased_random_states foo) = (FST x' IN (biased_high_states foo) /\
2484                                                     SND x' IN (biased_random_states foo))`
2485                        by (RW_TAC std_ss [] >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
2486                            >> POP_ORW >> RW_TAC std_ss []
2487                            >> RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND])
2488                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2489                >> `(\(x' :bool state # bool state). (if x' IN
2490                        (biased_high_states foo CROSS biased_random_states foo)
2491                        then 3 / 2 * (if FST x' IN biased_high_states foo /\ SND x' IN biased_random_states foo
2492                        then 1 / 16 else 0) else 0)) =
2493                    (\(x' :bool state # bool state). (if x' IN
2494                        (biased_high_states foo CROSS biased_random_states foo)
2495                        then (\x. 3 / 32) x' else 0))`
2496                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [IN_CROSS])
2497                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM, CARD_CROSS, FINITE_CROSS],
2498                Suff `PREIMAGE L {z} INTER (biased_high_states foo CROSS biased_low_states CROSS
2499                                                biased_random_states foo) = {}`
2500                >- RW_TAC bool_ss [REAL_SUM_IMAGE_THM]
2501                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, L_def, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING, biased_low_states_def]
2502                >> METIS_TAC []])
2503   >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2504   >> `!x y z.
2505        SIGMA
2506          (\s.
2507             (if SND (FST s) = (\s. s = STRCAT "coin" (toString 0)) then
2508                1 / 2 *
2509                (if
2510                   s IN
2511                   biased_high_states foo CROSS biased_low_states CROSS
2512                   biased_random_states foo
2513                 then
2514                   1 / 16
2515                 else
2516                   0)
2517              else
2518                3 / 2 *
2519                (if
2520                   s IN
2521                   biased_high_states foo CROSS biased_low_states CROSS
2522                   biased_random_states foo
2523                 then
2524                   1 / 16
2525                 else
2526                   0)))
2527          (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2528           PREIMAGE FST {(y,z)} INTER
2529           (biased_high_states foo CROSS biased_low_states CROSS
2530            biased_random_states foo)) =
2531        (if (z = (\s. s = STRCAT "coin" (toString 0))) /\ y IN (biased_high_states foo) then
2532                1/32
2533         else (if (z = (\s:string. F)) /\ y IN (biased_high_states foo) then
2534                3/32
2535         else 0)) *
2536        &(CARD {r | r IN biased_random_states foo /\
2537                            (insider_dcprog (SUC foo) ((y,z),r) = x)})`
2538        by (RW_TAC real_ss []
2539            >| [`(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2540                 PREIMAGE FST {(y,(\s. s = STRCAT "coin" (toString 0)))} INTER
2541                (biased_high_states foo CROSS biased_low_states CROSS
2542                 biased_random_states foo)) =
2543                 IMAGE (\r. ((y, (\s. s = STRCAT "coin" (toString 0))),r))
2544                        {r | r IN biased_random_states foo /\
2545                                (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)}`
2546                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2547                                    IN_CROSS, FST]
2548                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2549                     >- (Q.PAT_X_ASSUM `P = (y,(\s:string. s = STRCAT "coin" (toString 0)))` (ASSUME_TAC o GSYM)
2550                         >> Q.EXISTS_TAC `SND x'`
2551                         >> RW_TAC std_ss [PAIR])
2552                     >> RW_TAC set_ss [biased_low_states_def])
2553                >> POP_ORW
2554                >> `FINITE {r | r IN biased_random_states foo /\
2555                                        (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)}`
2556                   by ((MP_TAC o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
2557                       >> RW_TAC bool_ss []
2558                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
2559                >> `INJ (\r. ((y,(\s. s = STRCAT "coin" (toString 0))),r))
2560                        {r | r IN biased_random_states foo /\
2561                                        (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2562                        (IMAGE (\r. ((y,(\s. s = STRCAT "coin" (toString 0))),r))
2563                                {r | r IN biased_random_states foo /\
2564                                        (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)})`
2565                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
2566                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF] >> POP_ASSUM (K ALL_TAC)
2567                >> `!(x' :bool state).
2568                         (((y,(\s. s = STRCAT "coin" (toString 0))),x') IN
2569                                biased_high_states foo CROSS biased_low_states CROSS
2570                                biased_random_states foo) = (x' IN (biased_random_states foo))`
2571                        by RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND]
2572                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2573                >> `(\(x' :bool state). (if x' IN
2574                        {r |r IN biased_random_states foo /\
2575                                (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2576                        then 1 / 2 * (if x' IN biased_random_states foo
2577                        then 1 / 16 else 0) else 0)) =
2578                    (\(x' :bool state). (if x' IN
2579                        {r |r IN biased_random_states foo /\
2580                                (insider_dcprog (SUC foo) ((y,(\s. s = STRCAT "coin" (toString 0))),r) = x)}
2581                        then (\x. 1 / 32) x' else 0))`
2582                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [GSPECIFICATION])
2583                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM],
2584                `(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2585                 PREIMAGE FST {(y,(\s:string. F))} INTER
2586                (biased_high_states foo CROSS biased_low_states CROSS
2587                 biased_random_states foo)) =
2588                 IMAGE (\r. ((y, (\s:string. F)),r))
2589                        {r | r IN biased_random_states foo /\
2590                                (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)}`
2591                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2592                                    IN_CROSS, FST]
2593                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2594                     >- (Q.PAT_X_ASSUM `P = (y,(\s:string. F))` (ASSUME_TAC o GSYM)
2595                         >> Q.EXISTS_TAC `SND x'`
2596                         >> RW_TAC std_ss [PAIR])
2597                     >> RW_TAC set_ss [biased_low_states_def])
2598                >> POP_ORW
2599                >> `FINITE {r | r IN biased_random_states foo /\
2600                                        (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)}`
2601                   by ((MP_TAC o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
2602                       >> RW_TAC bool_ss []
2603                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
2604                >> `INJ (\r. ((y,(\s:string. F)),r))
2605                        {r | r IN biased_random_states foo /\
2606                                        (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)}
2607                        (IMAGE (\r. ((y,(\s:string. F)),r))
2608                                {r | r IN biased_random_states foo /\
2609                                        (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)})`
2610                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
2611                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF] >> POP_ASSUM (K ALL_TAC)
2612                >> FULL_SIMP_TAC bool_ss [DE_MORGAN_THM]
2613                >> `!(x' :bool state).
2614                         (((y,(\s. F)),x') IN
2615                        biased_high_states foo CROSS biased_low_states CROSS
2616                        biased_random_states foo) = (x' IN (biased_random_states foo))`
2617                        by RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND]
2618                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2619                >> `(\(x' :bool state). (if x' IN
2620                        {r | r IN biased_random_states foo /\
2621                                        (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)}
2622                        then 3 / 2 * (if x' IN biased_random_states foo
2623                        then 1 / 16 else 0) else 0)) =
2624                    (\(x' :bool state). (if x' IN
2625                        {r | r IN biased_random_states foo /\
2626                                        (insider_dcprog (SUC foo) ((y,(\s:string. F)),r) = x)}
2627                        then (\x. 3 / 32) x' else 0))`
2628                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [GSPECIFICATION])
2629                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM],
2630                RW_TAC bool_ss [GSYM INTER_ASSOC]
2631                >> Suff `PREIMAGE FST {(y,z)} INTER (biased_high_states foo CROSS biased_low_states CROSS
2632                                                biased_random_states foo) = {}`
2633                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
2634                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, FST, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING, biased_low_states_def]
2635                >> METIS_TAC [PAIR, FST,SND]])
2636   >> `!x y z. (PREIMAGE (\x. (insider_dcprog (SUC foo) x,FST x))
2637           {(x,y,z)} INTER
2638         (biased_high_states foo CROSS biased_low_states CROSS
2639          biased_random_states foo)) =
2640        (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
2641           PREIMAGE FST {(y,z)} INTER
2642           (biased_high_states foo CROSS biased_low_states CROSS
2643            biased_random_states foo))`
2644        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, IN_CROSS, IN_PREIMAGE, IN_SING, L_def])
2645   >> POP_ORW
2646   >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2647   >> `!x y z.
2648        SIGMA
2649             (\s.
2650                (if
2651                   SND (FST s) = (\s. s = STRCAT "coin" (toString 0))
2652                 then
2653                   1 / 2 *
2654                   (if
2655                      s IN
2656                      biased_high_states foo CROSS
2657                      biased_low_states CROSS biased_random_states foo
2658                    then
2659                      1 / 16
2660                    else
2661                      0)
2662                 else
2663                   3 / 2 *
2664                   (if
2665                      s IN
2666                      biased_high_states foo CROSS
2667                      biased_low_states CROSS biased_random_states foo
2668                    then
2669                      1 / 16
2670                    else
2671                      0)))
2672             (PREIMAGE FST {(y,z)} INTER
2673              (biased_high_states foo CROSS biased_low_states CROSS
2674               biased_random_states foo)) =
2675        (if (z = (\s. s = STRCAT "coin" (toString 0))) /\ y IN (biased_high_states foo) then
2676                1/32
2677         else (if (z = (\s:string. F)) /\ y IN (biased_high_states foo) then
2678                3/32
2679         else 0)) *
2680        &(CARD (biased_random_states foo))`
2681        by (RW_TAC real_ss []
2682            >| [`(PREIMAGE FST {(y,(\s. s = STRCAT "coin" (toString 0)))} INTER
2683                (biased_high_states foo CROSS biased_low_states CROSS
2684                 biased_random_states foo)) =
2685                 IMAGE (\r. ((y, (\s. s = STRCAT "coin" (toString 0))),r))
2686                        (biased_random_states foo)`
2687                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2688                                    IN_CROSS, FST]
2689                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2690                     >- (Q.PAT_X_ASSUM `P = (y,(\s:string. s = STRCAT "coin" (toString 0)))` (ASSUME_TAC o GSYM)
2691                         >> Q.EXISTS_TAC `SND x`
2692                         >> RW_TAC std_ss [PAIR])
2693                     >> RW_TAC set_ss [biased_low_states_def])
2694                >> POP_ORW
2695                >> `INJ (\r. ((y,(\s. s = STRCAT "coin" (toString 0))),r))
2696                        (biased_random_states foo)
2697                        (IMAGE (\r. ((y,(\s. s = STRCAT "coin" (toString 0))),r))
2698                                (biased_random_states foo))`
2699                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
2700                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF, FINITE_CROSS] >> POP_ASSUM (K ALL_TAC)
2701                >> `!(x' :bool state).
2702                        (((y,(\s. s = STRCAT "coin" (toString 0))),x') IN
2703                        biased_high_states foo CROSS biased_low_states CROSS
2704                        biased_random_states foo) = (x' IN (biased_random_states foo))`
2705                        by RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND]
2706                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2707                >> `(\x: bool state. (if x IN biased_random_states foo then 1 / 2 * (1 / 16) else 0)) =
2708                    (\(x' :bool state). (if x' IN biased_random_states foo then (\x. 1 / 32) x' else 0))`
2709                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [IN_CROSS])
2710                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3,
2711                                              REAL_MUL_COMM],
2712                `(PREIMAGE FST {(y,(\s:string. F))} INTER
2713                (biased_high_states foo CROSS biased_low_states CROSS
2714                 biased_random_states foo)) =
2715                 IMAGE (\r. ((y, (\s:string. F)),r))
2716                        (biased_random_states foo)`
2717                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
2718                                    IN_CROSS, FST]
2719                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
2720                     >- (Q.PAT_X_ASSUM `P = (y,(\s:string. F))` (ASSUME_TAC o GSYM)
2721                         >> Q.EXISTS_TAC `SND x`
2722                         >> RW_TAC std_ss [PAIR])
2723                     >> RW_TAC set_ss [biased_low_states_def])
2724                >> POP_ORW
2725                >> `INJ (\r. ((y,(\s:string. F)),r))
2726                        (biased_random_states foo)
2727                        (IMAGE (\r. ((y,(\s:string. F)),r))
2728                                (biased_random_states foo))`
2729                        by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
2730                >> RW_TAC std_ss [REAL_SUM_IMAGE_IMAGE, o_DEF, Once REAL_SUM_IMAGE_IN_IF, FINITE_CROSS] >> POP_ASSUM (K ALL_TAC)
2731                >> FULL_SIMP_TAC std_ss [DE_MORGAN_THM]
2732                >> `!(x' :bool state).
2733                         (((y,(\s:string.F)),x') IN
2734                        biased_high_states foo CROSS biased_low_states CROSS
2735                        biased_random_states foo) = (x' IN (biased_random_states foo))`
2736                        by RW_TAC set_ss [IN_CROSS, biased_low_states_def, FST, SND]
2737                >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2738                >> Q.UNABBREV_TAC `foo`
2739                >> `(\x:bool state. (if x IN biased_random_states 2 then 3 / 2 * (1 / 16) else 0)) =
2740                    (\(x' :bool state). (if x' IN biased_random_states 2
2741                        then (\x. 3 / 32) x' else 0))`
2742                        by (RW_TAC std_ss [FUN_EQ_THM] >> RW_TAC real_ss [IN_CROSS])
2743                >> POP_ORW >> RW_TAC real_ss [GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3, REAL_MUL_COMM],
2744                Suff `PREIMAGE FST {(y,z)} INTER (biased_high_states foo CROSS biased_low_states CROSS
2745                                                biased_random_states foo) = {}`
2746                >- RW_TAC bool_ss [REAL_SUM_IMAGE_THM]
2747                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, FST, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING, biased_low_states_def]
2748                >> METIS_TAC [PAIR, FST, SND]])
2749   >> RW_TAC std_ss [] >> POP_ASSUM (K ALL_TAC)
2750   >> Q.UNABBREV_TAC `foo`
2751   >> RW_TAC bool_ss [(COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)))``
2752                              (SIMP_CONV set_ss [biased_high_states_def, insider_high_states_set_def]
2753                               THENC (TRY_CONV insider_hl_dup_conv))
2754                                dc_dup_conv),
2755                        (COMPUTE_CARD ``(biased_random_states (SUC (SUC 0)))``
2756                              (SIMP_CONV bool_ss [biased_random_states_def] THENC
2757                                SIMP_CONV bool_ss [PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
2758                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
2759                                                     THENC (TRY_CONV dc_dup_conv)))))
2760                        dc_dup_conv)]
2761   >> Q.ABBREV_TAC `foo = SUC (SUC 0)`
2762   >> RW_TAC real_ss []
2763   >> `(IMAGE (insider_dcprog (SUC foo))
2764        (biased_high_states foo CROSS biased_low_states CROSS
2765         biased_random_states foo) CROSS biased_low_states) =
2766        (IMAGE (\out. (out, (\s.s = STRCAT "coin" (toString 0)))) (IMAGE (insider_dcprog (SUC foo))
2767        (biased_high_states foo CROSS biased_low_states CROSS
2768         biased_random_states foo))) UNION
2769        (IMAGE (\out. (out, (\s:string.F))) (IMAGE (insider_dcprog (SUC foo))
2770        (biased_high_states foo CROSS biased_low_states CROSS
2771         biased_random_states foo)))`
2772        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [biased_low_states_def, IN_UNION, IN_IMAGE, IN_CROSS, IN_IMAGE, IN_INSERT, IN_SING]
2773            >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW
2774            >> RW_TAC bool_ss [PAIR_EQ]
2775            >> EQ_TAC >> RW_TAC std_ss [FST, SND] >> RW_TAC std_ss [] >> METIS_TAC [PAIR, FST,SND])
2776   >> POP_ORW
2777   >> `DISJOINT (IMAGE (\out. (out,(\s. s = STRCAT "coin" (toString 0))))
2778        (IMAGE (insider_dcprog (SUC foo))
2779           (biased_high_states foo CROSS biased_low_states CROSS
2780            biased_random_states foo)))
2781      (IMAGE (\out. (out,(\s. F)))
2782        (IMAGE (insider_dcprog (SUC foo))
2783           (biased_high_states foo CROSS biased_low_states CROSS
2784            biased_random_states foo)))`
2785        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [NOT_IN_EMPTY, IN_INTER, IN_IMAGE]
2786            >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW
2787            >> RW_TAC bool_ss [PAIR_EQ]
2788            >> METIS_TAC [])
2789   >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE, FINITE_CROSS]
2790   >> POP_ASSUM (K ALL_TAC)
2791   >> `!c:bool state. INJ (\out. (out,c))
2792            (IMAGE (insider_dcprog (SUC foo))
2793            (biased_high_states foo CROSS biased_low_states CROSS
2794             biased_random_states foo))
2795          (IMAGE (\out. (out,c))
2796            (IMAGE (insider_dcprog (SUC foo))
2797            (biased_high_states foo CROSS biased_low_states CROSS
2798             biased_random_states foo)))`
2799        by RW_TAC std_ss [INJ_DEF, IN_IMAGE]
2800   >> RW_TAC real_ss [IMAGE_FINITE, REAL_SUM_IMAGE_IMAGE, FINITE_CROSS, o_DEF]
2801   >- METIS_TAC []
2802   >> POP_ASSUM MP_TAC >> POP_ASSUM (K ALL_TAC) >> STRIP_TAC
2803   >> `(IMAGE (insider_dcprog (SUC foo))
2804       (biased_high_states foo CROSS biased_low_states CROSS
2805        biased_random_states foo) CROSS
2806     (biased_high_states foo CROSS biased_low_states)) =
2807        (IMAGE (\(out,h). (out, (h,(\s.s = STRCAT "coin" (toString 0))))) ((IMAGE (insider_dcprog (SUC foo))
2808        (biased_high_states foo CROSS biased_low_states CROSS
2809         biased_random_states foo)) CROSS (biased_high_states foo))) UNION
2810        (IMAGE (\(out,h). (out, (h,(\s:string.F)))) ((IMAGE (insider_dcprog (SUC foo))
2811        (biased_high_states foo CROSS biased_low_states CROSS
2812         biased_random_states foo)) CROSS (biased_high_states foo)))`
2813        by ((ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [biased_low_states_def, IN_UNION, IN_IMAGE, IN_CROSS, IN_IMAGE, IN_INSERT, IN_SING]
2814            >> EQ_TAC >> RW_TAC std_ss [FST, SND] >> RW_TAC std_ss [])
2815            >| [DISJ1_TAC >> Q.EXISTS_TAC `(insider_dcprog (SUC foo) x', FST (SND x))`
2816                >> RW_TAC std_ss []
2817                >- (`x = (FST x, (FST(SND x),SND(SND x)))` by METIS_TAC [PAIR]
2818                    >> POP_ORW >> RW_TAC bool_ss [PAIR_EQ]
2819                    >> RW_TAC std_ss [FST, SND])
2820                >> METIS_TAC [],
2821                DISJ2_TAC >> Q.EXISTS_TAC `(insider_dcprog (SUC foo) x', FST (SND x))`
2822                >> RW_TAC std_ss []
2823                >- (`x = (FST x, (FST(SND x),SND(SND x)))` by METIS_TAC [PAIR]
2824                    >> POP_ORW >> RW_TAC bool_ss [PAIR_EQ]
2825                    >> RW_TAC std_ss [FST, SND])
2826                >> METIS_TAC [],
2827                DISJ1_TAC >> Q.EXISTS_TAC `(insider_dcprog (SUC foo) x', FST (SND x))`
2828                >> RW_TAC std_ss []
2829                >- (`x = (FST x, (FST(SND x),SND(SND x)))` by METIS_TAC [PAIR]
2830                    >> POP_ORW >> RW_TAC bool_ss [PAIR_EQ]
2831                    >> RW_TAC std_ss [FST, SND])
2832                >> METIS_TAC [],
2833                DISJ2_TAC >> Q.EXISTS_TAC `(insider_dcprog (SUC foo) x', FST (SND x))`
2834                >> RW_TAC std_ss []
2835                >- (`x = (FST x, (FST(SND x),SND(SND x)))` by METIS_TAC [PAIR]
2836                    >> POP_ORW >> RW_TAC bool_ss [PAIR_EQ]
2837                    >> RW_TAC std_ss [FST, SND])
2838                >> METIS_TAC [],
2839                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2840                >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
2841                >> METIS_TAC [],
2842                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2843                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2844                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2845                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2846                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2847                >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
2848                >> METIS_TAC [],
2849                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2850                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2851                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2852                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2853                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2854                >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
2855                >> METIS_TAC [],
2856                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2857                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2858                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2859                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2860                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2861                >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
2862                >> METIS_TAC [],
2863                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2864                >> POP_ORW >> RW_TAC std_ss [FST, SND],
2865                `x' = (FST x', SND x')` by RW_TAC std_ss [PAIR]
2866                >> POP_ORW >> RW_TAC std_ss [FST, SND]])
2867   >> POP_ORW
2868   >> `DISJOINT (IMAGE (\(out,h). (out, (h,(\s.s = STRCAT "coin" (toString 0))))) ((IMAGE (insider_dcprog (SUC foo))
2869        (biased_high_states foo CROSS biased_low_states CROSS
2870         biased_random_states foo)) CROSS (biased_high_states foo)))
2871        (IMAGE (\(out,h). (out, (h,(\s:string.F)))) ((IMAGE (insider_dcprog (SUC foo))
2872        (biased_high_states foo CROSS biased_low_states CROSS
2873         biased_random_states foo)) CROSS (biased_high_states foo)))`
2874        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [NOT_IN_EMPTY, IN_INTER, IN_IMAGE, IN_CROSS, biased_low_states_def,
2875                           IN_INSERT, IN_SING]
2876            >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW
2877            >> RW_TAC bool_ss []
2878            >> Cases_on `SND(SND x) = (\s. s = STRCAT "coin" (toString 0))`
2879            >- (DISJ2_TAC >> STRIP_TAC >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR] >> POP_ORW
2880                >> RW_TAC std_ss [] >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW
2881                >> RW_TAC bool_ss [PAIR_EQ] >> `SND x = (FST(SND x),SND(SND x))` by RW_TAC std_ss [PAIR] >> POP_ORW
2882                >> RW_TAC bool_ss [PAIR_EQ])
2883            >> DISJ1_TAC >> STRIP_TAC >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR] >> POP_ORW
2884            >> RW_TAC std_ss [] >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW
2885            >> RW_TAC bool_ss [PAIR_EQ] >> `SND x = (FST(SND x),SND(SND x))` by RW_TAC std_ss [PAIR] >> POP_ORW
2886            >> RW_TAC bool_ss [PAIR_EQ])
2887   >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE, FINITE_CROSS]
2888   >> POP_ASSUM (K ALL_TAC)
2889   >> `!c:bool state. INJ (\(out,h). (out,h,c))
2890            (IMAGE (insider_dcprog (SUC foo))
2891           (biased_high_states foo CROSS biased_low_states CROSS
2892            biased_random_states foo) CROSS biased_high_states foo)
2893          (IMAGE (\(out,h). (out,h,c))
2894            (IMAGE (insider_dcprog (SUC foo))
2895           (biased_high_states foo CROSS biased_low_states CROSS
2896            biased_random_states foo) CROSS biased_high_states foo))`
2897        by (RW_TAC std_ss [INJ_DEF, IN_IMAGE] >- METIS_TAC []
2898            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2899            >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
2900            >> POP_ORW >> RW_TAC std_ss []
2901            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
2902            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
2903            >> POP_ORW >> RW_TAC std_ss []
2904            >> METIS_TAC [PAIR, PAIR_EQ])
2905   >> RW_TAC real_ss [IMAGE_FINITE, REAL_SUM_IMAGE_IMAGE, FINITE_CROSS, o_DEF]
2906   >> POP_ASSUM (K ALL_TAC)
2907   >> `!x:bool state # bool state. x = (FST x, SND x)` by RW_TAC std_ss [PAIR]
2908   >> POP_ORW >> RW_TAC std_ss []
2909   >> Q.ABBREV_TAC `sumone = SIGMA
2910      (\out.
2911         1 / 32 *
2912         &
2913           (CARD
2914              {(h,r) |
2915               h IN biased_high_states foo /\
2916               r IN biased_random_states foo /\
2917               (insider_dcprog (SUC foo)
2918                  ((h,(\s. s = STRCAT "coin" (toString 0))),r) =
2919                out)}) *
2920         logr 2
2921           (1 / 32 *
2922            &
2923              (CARD
2924                 {(h,r) |
2925                  h IN biased_high_states foo /\
2926                  r IN biased_random_states foo /\
2927                  (insider_dcprog (SUC foo)
2928                     ((h,(\s. s = STRCAT "coin" (toString 0))),r) =
2929                   out)}) / (1 / 4)))
2930      (IMAGE (insider_dcprog (SUC foo))
2931         (biased_high_states foo CROSS biased_low_states CROSS
2932          biased_random_states foo)) +
2933    SIGMA
2934      (\out.
2935         3 / 32 *
2936         &
2937           (CARD
2938              {(h,r) |
2939               h IN biased_high_states foo /\
2940               r IN biased_random_states foo /\
2941               (insider_dcprog (SUC foo) ((h,(\s. F)),r) = out)}) *
2942         logr 2
2943           (3 / 32 *
2944            &
2945              (CARD
2946                 {(h,r) |
2947                  h IN biased_high_states foo /\
2948                  r IN biased_random_states foo /\
2949                  (insider_dcprog (SUC foo) ((h,(\s. F)),r) = out)}) /
2950            (3 / 4)))
2951      (IMAGE (insider_dcprog (SUC foo))
2952         (biased_high_states foo CROSS biased_low_states CROSS
2953          biased_random_states foo))`
2954   >> `FINITE (IMAGE (insider_dcprog (SUC foo))
2955      (biased_high_states foo CROSS biased_low_states CROSS
2956       biased_random_states foo) CROSS biased_high_states foo)`
2957        by RW_TAC std_ss [IMAGE_FINITE, FINITE_CROSS]
2958   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `(IMAGE (insider_dcprog (SUC foo))
2959      (biased_high_states foo CROSS biased_low_states CROSS
2960       biased_random_states foo) CROSS biased_high_states foo)`) REAL_SUM_IMAGE_IN_IF]
2961   >> `(\x.
2962      if
2963        x IN
2964        IMAGE (insider_dcprog (SUC foo))
2965          (biased_high_states foo CROSS biased_low_states CROSS
2966           biased_random_states foo) CROSS biased_high_states foo
2967      then
2968        (\x.
2969           (if SND x IN biased_high_states foo then 1 / 32 else 0) *
2970           &CARD
2971              {r |
2972               r IN biased_random_states foo /\
2973               (insider_dcprog (SUC foo)
2974                  ((SND x,(\s. s = STRCAT "coin" (toString 0))),r) =
2975                FST x)} *
2976           logr 2
2977             ((if SND x IN biased_high_states foo then 1 / 32 else 0) *
2978              &CARD
2979                 {r |
2980                  r IN biased_random_states foo /\
2981                  (insider_dcprog (SUC foo)
2982                     ((SND x,(\s. s = STRCAT "coin" (toString 0))),r) =
2983                   FST x)} /
2984              ((if SND x IN biased_high_states foo then 1 / 32 else 0) *
2985               4))) x
2986      else
2987        0) =
2988        (\x.
2989        (if
2990           x IN
2991           IMAGE (insider_dcprog (SUC foo))
2992             (biased_high_states foo CROSS biased_low_states CROSS
2993              biased_random_states foo) CROSS biased_high_states foo
2994         then
2995           (\x. (1 / 32) *
2996           &
2997             (CARD
2998                {r |
2999                 r IN biased_random_states foo /\
3000                 (insider_dcprog (SUC foo)
3001                    ((SND x,(\s. s = STRCAT "coin" (toString 0))),r) =
3002                  FST x)}) *
3003           logr 2
3004             ((1 / 32) *
3005              &
3006                (CARD
3007                   {r |
3008                    r IN biased_random_states foo /\
3009                    (insider_dcprog (SUC foo)
3010                       ((SND x,(\s. s = STRCAT "coin" (toString 0))),
3011                        r) =
3012                     FST x)}) /
3013              ((1 / 32) *
3014               4))) x
3015         else
3016           0))` by (RW_TAC std_ss [FUN_EQ_THM, IN_CROSS] >> RW_TAC std_ss [])
3017   >> POP_ORW
3018   >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_IN_IF, IMAGE_FINITE, FINITE_CROSS]
3019   >> Q.ABBREV_TAC `sumtwo = SIGMA
3020     (\x.
3021        1 / 32 *
3022        &
3023          (CARD
3024             {r |
3025              r IN biased_random_states foo /\
3026              (insider_dcprog (SUC foo)
3027                 ((SND x,(\s. s = STRCAT "coin" (toString 0))),r) =
3028               FST x)}) *
3029        logr 2
3030          (1 / 32 *
3031           &
3032             (CARD
3033                {r |
3034                 r IN biased_random_states foo /\
3035                 (insider_dcprog (SUC foo)
3036                    ((SND x,(\s. s = STRCAT "coin" (toString 0))),r) =
3037                  FST x)}) / (1 / 32 * 4)))
3038     (IMAGE (insider_dcprog (SUC foo))
3039        (biased_high_states foo CROSS biased_low_states CROSS
3040         biased_random_states foo) CROSS biased_high_states foo)`
3041   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `(IMAGE (insider_dcprog (SUC foo))
3042      (biased_high_states foo CROSS biased_low_states CROSS
3043       biased_random_states foo) CROSS biased_high_states foo)`) REAL_SUM_IMAGE_IN_IF]
3044
3045
3046   >> `(\x.
3047      if
3048        x IN
3049        IMAGE (insider_dcprog (SUC foo))
3050          (biased_high_states foo CROSS biased_low_states CROSS
3051           biased_random_states foo) CROSS biased_high_states foo
3052      then
3053        (\x.
3054           (if SND x IN biased_high_states foo then 3 / 32 else 0) *
3055           &CARD
3056              {r |
3057               r IN biased_random_states foo /\
3058               (insider_dcprog (SUC foo) ((SND x,(\s. F)),r) = FST x)} *
3059           logr 2
3060             ((if SND x IN biased_high_states foo then 3 / 32 else 0) *
3061              &CARD
3062                 {r |
3063                  r IN biased_random_states foo /\
3064                  (insider_dcprog (SUC foo) ((SND x,(\s. F)),r) =
3065                   FST x)} /
3066              ((if SND x IN biased_high_states foo then 3 / 32 else 0) *
3067               4))) x
3068      else
3069        0) =
3070        (\x.
3071        (if
3072           x IN
3073           IMAGE (insider_dcprog (SUC foo))
3074             (biased_high_states foo CROSS biased_low_states CROSS
3075              biased_random_states foo) CROSS biased_high_states foo
3076         then
3077           (\x. (3 / 32) *
3078           &
3079             (CARD
3080                {r |
3081                 r IN biased_random_states foo /\
3082                 (insider_dcprog (SUC foo) ((SND x,(\s. F)),r) =
3083                  FST x)}) *
3084           logr 2
3085             ((3 / 32) *
3086              &
3087                (CARD
3088                   {r |
3089                    r IN biased_random_states foo /\
3090                    (insider_dcprog (SUC foo) ((SND x,(\s. F)),r) =
3091                     FST x)}) /
3092              ((3 / 32) *
3093               4))) x
3094         else
3095           0))` by (RW_TAC std_ss [FUN_EQ_THM, IN_CROSS] >> RW_TAC std_ss [])
3096   >> POP_ORW
3097   >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_IN_IF, IMAGE_FINITE, FINITE_CROSS]
3098   >> Q.UNABBREV_TAC `sumtwo` >> Q.UNABBREV_TAC `sumone`
3099   >> RW_TAC real_ss [] >> Q.PAT_X_ASSUM `!f. f IN P` (K ALL_TAC)
3100   >> RW_TAC std_ss [GSYM lg_def]
3101   >> `!(out:bool state) (x:real) y c. x * c * lg (x * c / y) =
3102                (\c. x * c * lg (x * c / y)) c` by RW_TAC std_ss []
3103   >> POP_ORW
3104   >> Q.ABBREV_TAC `c1 = (\c. 1 / 32 * c * lg (1 / 32 * c / (1 / 4)))`
3105   >> Q.ABBREV_TAC `c2 = (\c. 3 / 32 * c * lg (3 / 32 * c / (3 / 4)))`
3106   >> Q.ABBREV_TAC `c3 = (\c. 1 / 32 * c * lg (1 / 32 * c / (1 / 8)))`
3107   >> Q.ABBREV_TAC `c4 = (\c. 3 / 32 * c * lg (3 / 32 * c / (3 / 8)))`
3108   >> `!(out:bool state) (b:bool state).
3109        &(CARD {(h,r) |
3110                  h IN biased_high_states foo /\
3111                  r IN biased_random_states foo /\
3112                  (insider_dcprog (SUC foo)
3113                     ((h,b),r) =
3114                   out)}) =
3115        REAL_SUM_IMAGE
3116        (\(h,r). if (insider_dcprog (SUC foo)
3117                        ((h,b),r) = out) then 1 else 0)
3118        (biased_high_states foo CROSS biased_random_states foo)`
3119        by (STRIP_TAC >> STRIP_TAC
3120            >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3121                                        (insider_dcprog (SUC foo) ((h,b),r) = out)}`
3122                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
3123                       >> RW_TAC bool_ss [FINITE_CROSS]
3124                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
3125                       >> POP_ASSUM MP_TAC
3126                       >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3127                       >> POP_ORW >> RW_TAC std_ss [])
3128            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
3129            >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3130                                        ~(insider_dcprog (SUC foo) ((h,b),r) = out)}`
3131                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
3132                       >> RW_TAC bool_ss [FINITE_CROSS]
3133                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
3134                       >> POP_ASSUM MP_TAC
3135                       >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3136                       >> POP_ORW >> RW_TAC std_ss [])
3137            >> `(biased_high_states foo CROSS biased_random_states foo) =
3138                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3139                                        (insider_dcprog (SUC foo) ((h,b),r) = out)} UNION
3140                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3141                                        ~(insider_dcprog (SUC foo) ((h,b),r) = out)}`
3142                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
3143                    >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
3144                    >- (Cases_on `insider_dcprog (SUC foo)
3145                                        ((FST x,b),SND x) = out`
3146                        >- (DISJ1_TAC >> Q.EXISTS_TAC `(FST x, SND x)` >> RW_TAC std_ss [PAIR_EQ])
3147                        >> DISJ2_TAC >> Q.EXISTS_TAC `(FST x, SND x)` >> RW_TAC std_ss [PAIR_EQ])
3148                    >> POP_ASSUM MP_TAC
3149                    >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
3150                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ])
3151            >> POP_ORW
3152            >> `DISJOINT {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3153                                        (insider_dcprog (SUC foo) ((h,b),r) = out)}
3154                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3155                                        ~(insider_dcprog (SUC foo) ((h,b),r) = out)}`
3156                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
3157                    >> Cases_on `insider_dcprog (SUC foo)
3158                                        ((FST x,b),SND x) = out`
3159                    >- (DISJ2_TAC >> STRIP_TAC
3160                        >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
3161                        >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
3162                    >> DISJ1_TAC >> STRIP_TAC
3163                    >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
3164                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
3165            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
3166            >> `SIGMA (\(h,r). (if insider_dcprog (SUC foo)
3167                        ((h,b),r) = out then 1 else 0))
3168                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3169                        ~(insider_dcprog (SUC foo)
3170                        ((h,b),r) = out)} = 0`
3171                by (RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
3172                    >> Suff `!x. (if x IN
3173                        {(h,r) |
3174                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
3175                        ~(insider_dcprog (SUC foo)
3176                        ((h,b),r) = out)} then
3177                        (\(h,r). (if insider_dcprog (SUC foo)
3178                        ((h,b),r) =
3179                        out then 1 else 0)) x else 0) = 0`
3180                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
3181                    >> RW_TAC std_ss [GSPECIFICATION]
3182                    >> POP_ASSUM MP_TAC >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
3183                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3184                    >> `x:bool state # bool state = (FST x, SND x)` by RW_TAC std_ss [PAIR]
3185                    >> POP_ORW >> RW_TAC std_ss [])
3186            >> POP_ORW >> RW_TAC real_ss []
3187            >> Suff `(\x. (if x IN {(h,r) |
3188                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
3189                        (insider_dcprog (SUC foo) ((h,b),r) = out)}
3190                        then 1 else 0)) =
3191                     (\x. (if x IN {(h,r) |
3192                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
3193                        (insider_dcprog (SUC foo) ((h,b),r) = out)}
3194                        then (\(h,r). (if insider_dcprog (SUC foo)
3195                                ((h,b),r) =
3196                        out then 1 else 0)) x else 0))`
3197             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3198                 >> METIS_TAC [])
3199             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION]
3200             >> POP_ASSUM MP_TAC >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
3201             >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3202             >> `x:bool state # bool state = (FST x, SND x)` by RW_TAC std_ss [PAIR]
3203             >> POP_ORW >> RW_TAC std_ss [])
3204   >> POP_ORW
3205   >> `!(x:bool state # bool state) (b:bool state).
3206        &(CARD {r | r IN biased_random_states foo /\
3207                  (insider_dcprog (SUC foo)
3208                     ((SND x,b),r) = FST x)}) =
3209        REAL_SUM_IMAGE
3210        (\r. if (insider_dcprog (SUC foo)
3211                        ((SND x,b),r) = FST x) then 1 else 0)
3212        (biased_random_states foo)`
3213        by (STRIP_TAC >> STRIP_TAC
3214            >> `FINITE {r | r IN biased_random_states foo /\
3215                            (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}`
3216                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
3217                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
3218            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
3219            >> `FINITE {r | r IN biased_random_states foo /\
3220                            ~(insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}`
3221                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
3222                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
3223            >>  Q.ABBREV_TAC `s = {r | r IN biased_random_states foo /\
3224                                (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}`
3225            >> `(biased_random_states foo) =
3226                {r | r IN biased_random_states foo /\
3227                            (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)} UNION
3228                {r | r IN biased_random_states foo /\
3229                            ~(insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}`
3230                by (Q.UNABBREV_TAC `s` >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION] >> DECIDE_TAC)
3231            >> POP_ORW
3232            >> Q.UNABBREV_TAC `s`
3233            >> `DISJOINT {r | r IN biased_random_states foo /\
3234                            (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}
3235                {r | r IN biased_random_states foo /\
3236                            ~(insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}`
3237                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
3238                    >> METIS_TAC [])
3239            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
3240            >> `SIGMA (\r. (if insider_dcprog (SUC foo) ((SND x,b),r) = FST x then 1 else 0))
3241                {r | r IN biased_random_states foo /\
3242                            ~(insider_dcprog (SUC foo) ((SND x,b),r) = FST x)} = 0`
3243                by (RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
3244                    >> Suff `!r. (if r IN
3245                        {r | r IN biased_random_states foo /\
3246                            ~(insider_dcprog (SUC foo) ((SND x,b),r) = FST x)} then
3247                        (if insider_dcprog (SUC foo)
3248                        ((SND x,b),r) =
3249                        FST x then 1 else 0) else 0) = 0`
3250                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
3251                    >> RW_TAC std_ss [GSPECIFICATION])
3252            >> POP_ORW >> RW_TAC real_ss []
3253            >> Suff `(\x'. (if x' IN {r | r IN biased_random_states foo /\
3254                            (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}
3255                        then 1 else 0)) =
3256                     (\x'. (if x' IN {r | r IN biased_random_states foo /\
3257                            (insider_dcprog (SUC foo) ((SND x,b),r) = FST x)}
3258                        then (\r. (if insider_dcprog (SUC foo) ((SND x,b),r) = FST x then 1 else 0)) x' else 0))`
3259             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3260                 >> METIS_TAC [])
3261             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION])
3262   >> POP_ORW
3263   >> Q.UNABBREV_TAC `foo`
3264   >> Q.ABBREV_TAC `s1 = (IMAGE (insider_dcprog (SUC (SUC (SUC 0))))
3265         (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3266          biased_random_states (SUC (SUC 0))))`
3267   >> Q.ABBREV_TAC `s2 = (s1 CROSS biased_high_states (SUC (SUC 0)))`
3268   >> RW_TAC bool_ss [biased_high_states_def, biased_random_states_def, insider_high_states_set_def,
3269                      insider_dcprog_def, compute_result_alt, insider_set_announcements_alt, XOR_announces_def,
3270                      H_def, L_def, PAIR, ETA_THM, xor_def, STRCAT_toString_inj]
3271   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3272                                                   L_def, SND]
3273                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3274                                                     THENC (TRY_CONV insider_hl_dup_conv)
3275                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
3276   >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
3277                         THENC ((ONCE_FIND_CONV ``x DELETE y``
3278                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3279                                              THENC EVAL
3280                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3281                                              THENC EVAL
3282                                              THENC (REPEATC (T_F_UNCHANGED_CONV
3283                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3284                                                         THENC EVAL)))))))
3285                         THENC SIMP_CONV arith_ss []))
3286   >> Q.UNABBREV_TAC `s2` >> Q.UNABBREV_TAC `s1`
3287   >> RW_TAC bool_ss [biased_high_states_def, biased_low_states_def, biased_random_states_def, insider_high_states_set_def,
3288                      insider_dcprog_def, compute_result_alt, insider_set_announcements_alt, XOR_announces_def,
3289                      H_def, L_def, PAIR, ETA_THM, xor_def, STRCAT_toString_inj]
3290   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3291                                                   L_def, SND]
3292                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3293                                                     THENC (TRY_CONV insider_hl_dup_conv)
3294                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
3295   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3296                                                   L_def, SND]
3297                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3298                                                     THENC (TRY_CONV insider_hl_dup_conv)
3299                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
3300   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3301                                                   L_def, SND]
3302                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3303                                                     THENC (TRY_CONV insider_hl_dup_conv)
3304                                                     THENC (TRY_CONV insider_hl_dup_conv)
3305                                                     THENC (TRY_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3306                                              THENC EVAL
3307                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3308                                              THENC EVAL
3309                                              THENC (REPEATC (T_F_UNCHANGED_CONV
3310                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3311                                                         THENC EVAL)))))))))
3312  >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
3313                         THENC ((ONCE_FIND_CONV ``x DELETE y``
3314                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3315                                              THENC EVAL
3316                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3317                                              THENC EVAL
3318                                              THENC (REPEATC (T_F_UNCHANGED_CONV
3319                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3320                                                         THENC EVAL)))))))
3321                         THENC SIMP_CONV arith_ss []))
3322   >> CONV_TAC (ONCE_FIND_CONV ``if (x=y) then (1:real) else 0`` (RATOR_CONV (RATOR_CONV (RAND_CONV
3323        (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3324         THENC EVAL
3325         THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3326         THENC EVAL
3327         THENC (REPEATC (T_F_UNCHANGED_CONV
3328         (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3329         THENC EVAL))))))
3330                                                                  THENC SIMP_CONV bool_ss []))
3331   >> RW_TAC real_ss []
3332   >> Q.UNABBREV_TAC `c4` >> Q.UNABBREV_TAC `c3` >> Q.UNABBREV_TAC `c2` >> Q.UNABBREV_TAC `c1`
3333   >> RW_TAC real_ss []
3334   >> `lg(1/4) = ~2` by (RW_TAC real_ss [GSYM REAL_INV_1OVER, lg_inv] >> `4 = 2 pow 2` by RW_TAC real_ss [pow] >> POP_ORW >> RW_TAC std_ss [lg_pow])
3335   >> POP_ORW
3336   >> RW_TAC real_ss []);
3337
3338(* ************************************************************************* *)(* Biasing a hidden coin creates a partial leak                              *)(* ************************************************************************* *)
3339
3340val biased_dist2_def = Define
3341   `biased_dist2 high low random =
3342        (\s. if R s (STRCAT "coin" (toString (SUC 0))) then
3343                (1 / 2) * (unif_prog_dist high low random s)
3344             else
3345                (3 / 2) * (unif_prog_dist high low random s))`;
3346
3347val biased_dc_prog_space2_def = Define   `biased_dc_prog_space2 (SUC(SUC n)) =
3348        (((biased_high_states (SUC(SUC n))) CROSS biased_low_states) CROSS (biased_random_states (SUC (SUC n))),
3349         POW (((biased_high_states (SUC(SUC n))) CROSS biased_low_states) CROSS (biased_random_states (SUC (SUC n)))),
3350         (\s. SIGMA (biased_dist2 (biased_high_states (SUC(SUC n))) biased_low_states (biased_random_states (SUC (SUC n)))) s))`;
3351val prob_space_biased_dc_prog_space23 = store_thm
3352  ("prob_space_biased_dc_prog_space23",
3353   ``prob_space (biased_dc_prog_space2 (SUC(SUC 0)))``,
3354   (RW_TAC bool_ss [prob_space_def]
3355    >> `FINITE ((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3356                biased_random_states (SUC (SUC 0))))`
3357        by RW_TAC set_ss [biased_high_states_def, biased_low_states_def,
3358                             biased_random_states_def, insider_high_states_set_def])
3359   >- (MATCH_MP_TAC finite_additivity_sufficient_for_finite_spaces2
3360       >> RW_TAC bool_ss [biased_dc_prog_space2_def, m_space_def, measurable_sets_def, POW_SIGMA_ALGEBRA,
3361                          PSPACE, measure_def, positive_def, REAL_SUM_IMAGE_THM, IN_POW,
3362                          additive_def, IN_FUNSET, IN_UNIV]
3363       >- (MATCH_MP_TAC REAL_SUM_IMAGE_POS
3364           >> CONJ_TAC
3365           >- METIS_TAC [SUBSET_FINITE]
3366           >> RW_TAC real_ss [biased_dist2_def, unif_prog_dist_def, GSYM REAL_INV_1OVER, REAL_LE_INV_EQ, REAL_LE_MUL])
3367       >> METIS_TAC [SUBSET_FINITE, REAL_SUM_IMAGE_DISJOINT_UNION])
3368   >> RW_TAC bool_ss [biased_dc_prog_space2_def, m_space_def, measurable_sets_def,
3369                      PSPACE, measure_def, biased_dist2_def]
3370   >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3371        biased_random_states (SUC (SUC 0))) =
3372        (((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3373         biased_random_states (SUC (SUC 0))))INTER
3374         {s| R s (STRCAT "coin" (toString (SUC 0)))}) UNION
3375        (((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3376        biased_random_states (SUC (SUC 0))))INTER{s| ~R s (STRCAT "coin" (toString (SUC 0)))})`
3377        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, IN_INTER, GSPECIFICATION] >> DECIDE_TAC)
3378   >> POP_ORW
3379   >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3380        biased_random_states (SUC (SUC 0)) INTER
3381        {s | R s (STRCAT "coin" (toString (SUC 0)))}) /\
3382        FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3383        biased_random_states (SUC (SUC 0)) INTER
3384        {s | ~R s (STRCAT "coin" (toString (SUC 0)))})`
3385        by RW_TAC std_ss [INTER_FINITE]
3386   >> `DISJOINT (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3387        biased_random_states (SUC (SUC 0)) INTER
3388        {s | R s (STRCAT "coin" (toString (SUC 0)))})
3389        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3390        biased_random_states (SUC (SUC 0)) INTER
3391        {s | ~R s (STRCAT "coin" (toString (SUC 0)))})`
3392        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
3393            >> DECIDE_TAC)
3394   >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION]
3395   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3396        biased_random_states (SUC (SUC 0)) INTER
3397        {s | R s (STRCAT "coin" (toString (SUC 0)))}`) REAL_SUM_IMAGE_IN_IF]
3398   >> RW_TAC bool_ss [unif_prog_dist_def]
3399   >> `(\x.
3400     if
3401       x IN
3402       biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3403       biased_random_states (SUC (SUC 0)) INTER
3404       {s | R s (STRCAT "coin" (toString (SUC 0)))}
3405     then
3406       if R x (STRCAT "coin" (toString (SUC 0))) then
3407         1 / 2 *
3408         if
3409           x IN
3410           biased_high_states (SUC (SUC 0)) CROSS
3411           biased_low_states CROSS biased_random_states (SUC (SUC 0))
3412         then
3413           1 /
3414           &CARD
3415              (biased_high_states (SUC (SUC 0)) CROSS
3416               biased_low_states CROSS
3417               biased_random_states (SUC (SUC 0)))
3418         else
3419           0
3420       else
3421         3 / 2 *
3422         if
3423           x IN
3424           biased_high_states (SUC (SUC 0)) CROSS
3425           biased_low_states CROSS biased_random_states (SUC (SUC 0))
3426         then
3427           1 /
3428           &CARD
3429              (biased_high_states (SUC (SUC 0)) CROSS
3430               biased_low_states CROSS
3431               biased_random_states (SUC (SUC 0)))
3432         else
3433           0
3434     else
3435       0) =
3436        (\s. (if s IN
3437        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3438        biased_random_states (SUC (SUC 0)) INTER
3439        {s | R s (STRCAT "coin" (toString (SUC 0)))}
3440      then
3441        (\s. 1 / 2 * (1 /
3442              &
3443                (CARD
3444                   (biased_high_states (SUC (SUC 0)) CROSS
3445                    biased_low_states CROSS
3446                    biased_random_states (SUC (SUC 0)))))) s
3447      else 0))`
3448        by (RW_TAC bool_ss [Once FUN_EQ_THM]
3449            >> RW_TAC arith_ss [IN_INTER, GSPECIFICATION])
3450   >> POP_ORW
3451   >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3452   >> RW_TAC bool_ss [REAL_SUM_IMAGE_FINITE_CONST3]
3453   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3454        biased_random_states (SUC (SUC 0)) INTER
3455        {s |~ R s (STRCAT "coin" (toString (SUC 0)))}`) REAL_SUM_IMAGE_IN_IF]
3456   >> RW_TAC bool_ss [unif_prog_dist_def]
3457   >> `(\s.
3458     (if
3459        s IN
3460        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3461        biased_random_states (SUC (SUC 0)) INTER
3462        {s | ~R s (STRCAT "coin" (toString (SUC 0)))}
3463      then
3464        (if R s (STRCAT "coin" (toString (SUC 0))) then
3465           1 / 2 *
3466           (if
3467              s IN
3468              biased_high_states (SUC (SUC 0)) CROSS
3469              biased_low_states CROSS biased_random_states (SUC (SUC 0))
3470            then
3471              1 /
3472              &
3473                (CARD
3474                   (biased_high_states (SUC (SUC 0)) CROSS
3475                    biased_low_states CROSS
3476                    biased_random_states (SUC (SUC 0))))
3477            else
3478              0)
3479         else
3480           3/2 *
3481           (if
3482              s IN
3483              biased_high_states (SUC (SUC 0)) CROSS
3484              biased_low_states CROSS biased_random_states (SUC (SUC 0))
3485            then
3486              1 /
3487              &
3488                (CARD
3489                   (biased_high_states (SUC (SUC 0)) CROSS
3490                    biased_low_states CROSS
3491                    biased_random_states (SUC (SUC 0))))
3492            else
3493              0))
3494      else
3495        0)) =
3496        (\s. (if s IN
3497        biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3498        biased_random_states (SUC (SUC 0)) INTER
3499        {s | ~R s (STRCAT "coin" (toString (SUC 0)))}
3500      then
3501        (\s. 3/2 * (1 /
3502              &
3503                (CARD
3504                   (biased_high_states (SUC (SUC 0)) CROSS
3505                    biased_low_states CROSS
3506                    biased_random_states (SUC (SUC 0)))))) s
3507      else 0))`
3508        by (RW_TAC bool_ss [Once FUN_EQ_THM]
3509            >> RW_TAC bool_ss [IN_INTER, GSPECIFICATION])
3510   >> POP_ORW
3511   >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3512   >> RW_TAC bool_ss [REAL_SUM_IMAGE_FINITE_CONST3]
3513   >> `!b. & (CARD
3514     (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3515      biased_random_states (SUC (SUC 0)) INTER
3516      {s | R s b})) =
3517        REAL_SUM_IMAGE
3518        (\s. if R s b then 1 else 0)
3519        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3520      biased_random_states (SUC (SUC 0)))`
3521        by (STRIP_TAC
3522            >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3523                biased_random_states (SUC (SUC 0)) INTER
3524                {s | R s b})`
3525                   by ((MP_TAC o Q.ISPEC `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3526                        biased_random_states (SUC (SUC 0)))`) SUBSET_FINITE
3527                       >> RW_TAC std_ss [FINITE_CROSS] >> POP_ASSUM MATCH_MP_TAC
3528                       >> RW_TAC std_ss [SUBSET_DEF, IN_INTER])
3529            >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3530                biased_random_states (SUC (SUC 0)) INTER
3531                {s | ~R s b})`
3532                   by ((MP_TAC o Q.ISPEC `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3533                        biased_random_states (SUC (SUC 0)))`) SUBSET_FINITE
3534                       >> RW_TAC std_ss [FINITE_CROSS] >> POP_ASSUM MATCH_MP_TAC
3535                       >> RW_TAC std_ss [SUBSET_DEF, IN_INTER])
3536            >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
3537            >>  Q.ABBREV_TAC `s = (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3538                                biased_random_states (SUC (SUC 0)) INTER {s | R s b})`
3539            >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3540                biased_random_states (SUC (SUC 0))) =
3541                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3542                biased_random_states (SUC (SUC 0)) INTER {s | R s b}) UNION
3543                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3544                biased_random_states (SUC (SUC 0)) INTER {s | ~R s b})`
3545                by (Q.UNABBREV_TAC `s` >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, IN_INTER, GSPECIFICATION] >> DECIDE_TAC)
3546            >> POP_ORW
3547            >> Q.UNABBREV_TAC `s`
3548            >> `DISJOINT (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3549                biased_random_states (SUC (SUC 0)) INTER {s | R s b})
3550                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3551                biased_random_states (SUC (SUC 0)) INTER {s | ~R s b})`
3552                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
3553                    >> METIS_TAC [])
3554            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
3555            >> `SIGMA (\s. (if R s b then 1 else 0))
3556                        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3557                        biased_random_states (SUC (SUC 0)) INTER {s | ~R s b}) = 0`
3558                by (RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
3559                    >> RW_TAC bool_ss [IN_INTER, GSPECIFICATION, REAL_SUM_IMAGE_0])
3560            >> POP_ORW >> RW_TAC real_ss []
3561            >> FULL_SIMP_TAC std_ss []
3562            >> Suff `(\x. (if x IN biased_high_states 2 CROSS biased_low_states CROSS
3563                        biased_random_states 2 INTER {s | R s b}
3564                        then 1 else 0)) =
3565                     (\x. (if x IN biased_high_states 2 CROSS biased_low_states CROSS
3566                        biased_random_states 2 INTER {s | R s b}
3567                        then (\s. if R s b then 1 else 0) x else 0))`
3568             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3569                 >> METIS_TAC [])
3570             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [IN_INTER,GSPECIFICATION])
3571   >> POP_ORW
3572   >> `!b. & (CARD
3573     (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3574      biased_random_states (SUC (SUC 0)) INTER
3575      {s | ~R s b})) =
3576        REAL_SUM_IMAGE
3577        (\s. if ~R s b then 1 else 0)
3578        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3579      biased_random_states (SUC (SUC 0)))`
3580        by (STRIP_TAC
3581            >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3582                biased_random_states (SUC (SUC 0)) INTER
3583                {s | R s b})`
3584                   by ((MP_TAC o Q.ISPEC `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3585                        biased_random_states (SUC (SUC 0)))`) SUBSET_FINITE
3586                       >> RW_TAC std_ss [FINITE_CROSS] >> POP_ASSUM MATCH_MP_TAC
3587                       >> RW_TAC std_ss [SUBSET_DEF, IN_INTER])
3588            >> `FINITE (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3589                biased_random_states (SUC (SUC 0)) INTER
3590                {s | ~R s b})`
3591                   by ((MP_TAC o Q.ISPEC `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3592                        biased_random_states (SUC (SUC 0)))`) SUBSET_FINITE
3593                       >> RW_TAC std_ss [FINITE_CROSS] >> POP_ASSUM MATCH_MP_TAC
3594                       >> RW_TAC std_ss [SUBSET_DEF, IN_INTER])
3595            >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
3596            >>  Q.ABBREV_TAC `s = (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3597                                biased_random_states (SUC (SUC 0)) INTER {s | ~R s b})`
3598            >> `(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3599                biased_random_states (SUC (SUC 0))) =
3600                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3601                biased_random_states (SUC (SUC 0)) INTER {s | R s b}) UNION
3602                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3603                biased_random_states (SUC (SUC 0)) INTER {s | ~R s b})`
3604                by (Q.UNABBREV_TAC `s` >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, IN_INTER, GSPECIFICATION] >> DECIDE_TAC)
3605            >> POP_ORW
3606            >> Q.UNABBREV_TAC `s`
3607            >> `DISJOINT (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3608                biased_random_states (SUC (SUC 0)) INTER {s | R s b})
3609                (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3610                biased_random_states (SUC (SUC 0)) INTER {s | ~R s b})`
3611                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
3612                    >> METIS_TAC [])
3613            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
3614            >> `SIGMA (\s. (if ~R s b then 1 else 0))
3615                        (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3616                        biased_random_states (SUC (SUC 0)) INTER {s | R s b}) = 0`
3617                by (RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
3618                    >> RW_TAC bool_ss [IN_INTER, GSPECIFICATION, REAL_SUM_IMAGE_0])
3619            >> POP_ORW >> RW_TAC real_ss []
3620            >> FULL_SIMP_TAC std_ss []
3621            >> Suff `(\x. (if x IN biased_high_states 2 CROSS biased_low_states CROSS
3622                        biased_random_states 2 INTER {s | ~R s b}
3623                        then 1 else 0)) =
3624                     (\x. (if x IN biased_high_states 2 CROSS biased_low_states CROSS
3625                        biased_random_states 2 INTER {s | ~R s b}
3626                        then (\s. if ~R s b then 1 else 0) x else 0))`
3627             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
3628                 >> METIS_TAC [])
3629             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [IN_INTER,GSPECIFICATION])
3630   >> POP_ORW
3631   >> RW_TAC bool_ss [          (COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3632                                biased_random_states (SUC (SUC 0)))``
3633                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
3634                                biased_random_states_def, insider_high_states_set_def] THENC
3635                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
3636                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3637                                                     THENC (TRY_CONV insider_hl_dup_conv)
3638                                                     THENC (TRY_CONV insider_hl_dup_conv))))
3639                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
3640         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3641                                                     THENC (TRY_CONV insider_hl_dup_conv)
3642                                                     THENC (TRY_CONV insider_hl_dup_conv)
3643                                                     THENC (TRY_CONV dc_dup_conv)))))
3644                dc_dup_conv)]
3645   >> RW_TAC bool_ss [biased_high_states_def, biased_low_states_def, biased_random_states_def, insider_high_states_set_def,
3646                      insider_dcprog_def, compute_result_alt, insider_set_announcements_alt, XOR_announces_def,
3647                      H_def, L_def, R_def, PAIR, ETA_THM, xor_def, STRCAT_toString_inj]
3648   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3649                                                   H_def, L_def, R_def, SND]
3650                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3651                                                     THENC (TRY_CONV insider_hl_dup_conv)
3652                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
3653   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3654                                                   H_def, L_def, R_def, SND]
3655                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3656                                                     THENC (TRY_CONV insider_hl_dup_conv)
3657                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
3658   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
3659                                                   H_def, L_def, R_def, SND]
3660                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3661                                                     THENC (TRY_CONV insider_hl_dup_conv)
3662                                                     THENC (TRY_CONV insider_hl_dup_conv)
3663                                                     THENC (TRY_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3664                                              THENC EVAL
3665                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3666                                              THENC EVAL
3667                                              THENC (REPEATC (T_F_UNCHANGED_CONV
3668                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3669                                                         THENC EVAL)))))))))
3670  >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
3671                         THENC ((ONCE_FIND_CONV ``x DELETE y``
3672                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
3673                                              THENC EVAL
3674                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
3675                                              THENC EVAL
3676                                              THENC (REPEATC (T_F_UNCHANGED_CONV
3677                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
3678                                                         THENC EVAL)))))))
3679                         THENC SIMP_CONV arith_ss []))
3680   >> RW_TAC real_ss [STRCAT_toString_inj]);
3681
3682(* ------------------------------------------------------------------------- *)
3683
3684val biased_dc3_leakage_result2 = store_thm
3685  ("biased_dc3_leakage_result2",
3686   ``(leakage (biased_dc_prog_space2 (SUC (SUC 0))) (insider_dcprog (SUC(SUC(SUC 0)))) = 3 / 4 * lg 3 - 1)``,
3687   RW_TAC bool_ss [leakage_def]
3688   >> Q.ABBREV_TAC `v = 3 / 4 * lg 3 - 1`
3689   >> `FINITE ((biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3690                biased_random_states (SUC (SUC 0))))`
3691        by RW_TAC set_ss [biased_high_states_def, biased_low_states_def,
3692                             biased_random_states_def, insider_high_states_set_def]
3693   >> `FINITE (biased_high_states (SUC (SUC 0)))`
3694        by RW_TAC set_ss [biased_high_states_def, insider_high_states_set_def]
3695   >> `FINITE (biased_random_states (SUC (SUC 0)))`
3696        by RW_TAC set_ss [biased_random_states_def]
3697   >> `FINITE (biased_low_states)`
3698        by RW_TAC set_ss [biased_low_states_def]
3699   >> (MP_TAC o
3700       Q.SPECL [`2`,`(biased_dc_prog_space2 (SUC (SUC 0)))`, `(insider_dcprog (SUC (SUC (SUC 0))))`, `H`,`L`] o
3701       INST_TYPE [alpha |-> ``:(bool, bool, bool) prog_state``, beta |-> ``:bool state``, gamma |-> ``:bool state``,
3702                  delta |-> ``:bool state``])
3703        finite_conditional_mutual_information_reduce
3704   >> (MP_TAC o INST_TYPE [beta |-> ``:bool state``] o Q.ISPEC `(biased_dc_prog_space2 (SUC (SUC 0)))`)
3705        MEASURABLE_POW_TO_POW_IMAGE
3706   >> RW_TAC bool_ss [random_variable_def, prob_space_biased_dc_prog_space23, biased_dc_prog_space2_def, PSPACE, EVENTS,
3707                      prob_space_def, measurable_sets_def, m_space_def,
3708                      REWRITE_RULE [prob_space_def] prob_space_biased_dc_prog_space23]
3709   >> POP_ASSUM (K ALL_TAC)
3710   >> RW_TAC bool_ss [biased_dist2_def, unif_prog_dist_def,
3711        (COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
3712                                biased_random_states (SUC (SUC 0)))``
3713                              (SIMP_CONV bool_ss [biased_high_states_def, biased_low_states_def,
3714                                biased_random_states_def, insider_high_states_set_def] THENC
3715                                SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
3716                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3717                                                     THENC (TRY_CONV insider_hl_dup_conv)
3718                                                     THENC (TRY_CONV insider_hl_dup_conv))))
3719                THENC SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY]
3720         THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
3721                                                     THENC (TRY_CONV insider_hl_dup_conv)
3722                                                     THENC (TRY_CONV insider_hl_dup_conv)
3723                                                     THENC (TRY_CONV dc_dup_conv)))))
3724                dc_dup_conv)]
3725   >> Q.ABBREV_TAC `foo = (SUC (SUC 0))`
3726   >> RW_TAC std_ss [joint_distribution_def, L_def, H_def, R_def, PAIR, ETA_THM, FST, SND, PROB, PSPACE, distribution_def]
3727   >> `IMAGE L
3728        (biased_high_states foo CROSS biased_low_states CROSS
3729         biased_random_states foo) =
3730        biased_low_states`
3731        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, L_def]
3732            >> EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []
3733            >> Q.UNABBREV_TAC `foo`
3734            >> Q.EXISTS_TAC `(((\s:string. s = STRCAT "pays" (toString (SUC (SUC 0)))),x),(\s:string. F))`
3735            >> RW_TAC bool_ss [SND, FST, biased_high_states_def, insider_high_states_set_def, IN_INSERT]
3736            >> Suff `!n. (\s:string. F) IN biased_random_states n` >- RW_TAC std_ss []
3737            >> REPEAT (POP_ASSUM (K ALL_TAC))
3738            >> Induct >> RW_TAC bool_ss [biased_random_states_def, IN_SING, IN_IMAGE, IN_UNION]
3739            >> DISJ2_TAC >> Q.EXISTS_TAC `(\s:string.F)` >> RW_TAC std_ss [])
3740   >> POP_ORW
3741   >> `(IMAGE (insider_dcprog (SUC foo))
3742        (biased_high_states foo CROSS biased_low_states CROSS
3743         biased_random_states foo) CROSS
3744      IMAGE FST
3745        (biased_high_states foo CROSS biased_low_states CROSS
3746         biased_random_states foo)) =
3747        (IMAGE (insider_dcprog (SUC foo))
3748        (biased_high_states foo CROSS biased_low_states CROSS
3749         biased_random_states foo) CROSS
3750      (biased_high_states foo CROSS biased_low_states))`
3751        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_IMAGE, IN_CROSS, FST]
3752            >> EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []
3753            >> Q.EXISTS_TAC `(SND x, SND x')`
3754            >> RW_TAC bool_ss [SND, FST])
3755   >> POP_ORW
3756   >> `!x z.
3757        SIGMA
3758          (\s.
3759             (if SND s (STRCAT "coin" (toString 1)) then
3760                1 / 2 *
3761                (if
3762                   s IN
3763                   biased_high_states foo CROSS biased_low_states CROSS
3764                   biased_random_states foo
3765                 then
3766                   1 / 16
3767                 else
3768                   0)
3769              else
3770                3 / 2 *
3771                (if
3772                   s IN
3773                   biased_high_states foo CROSS biased_low_states CROSS
3774                   biased_random_states foo
3775                 then
3776                   1 / 16
3777                 else
3778                   0)))
3779          (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
3780           PREIMAGE L {z} INTER
3781           (biased_high_states foo CROSS biased_low_states CROSS
3782            biased_random_states foo)) =
3783        (if z IN biased_low_states then
3784                1/32 * &(CARD {(h,r)| h IN biased_high_states foo /\ r IN biased_random_states foo /\
3785                       r (STRCAT "coin" (toString 1)) /\
3786                            (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
3787                3/32 * &(CARD {(h,r)| h IN biased_high_states foo /\ r IN biased_random_states foo /\
3788                       ~ r (STRCAT "coin" (toString 1)) /\
3789                            (insider_dcprog (SUC foo) ((h,z),r) = x)}) else 0)`
3790        by (REVERSE (RW_TAC real_ss [])
3791            >- (RW_TAC bool_ss [GSYM INTER_ASSOC]
3792                >> Suff `PREIMAGE L {z} INTER (biased_high_states foo CROSS biased_low_states CROSS
3793                                                biased_random_states foo) = {}`
3794                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
3795                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, L_def, IN_CROSS, IN_INTER, IN_PREIMAGE]
3796                >> METIS_TAC [])
3797            >> `(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
3798                 PREIMAGE L {z} INTER
3799                (biased_high_states foo CROSS biased_low_states CROSS
3800                 biased_random_states foo)) =
3801                 (IMAGE (\(h,r). ((h, z),r))
3802                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3803                                r (STRCAT "coin" (toString 1)) /\
3804                                (insider_dcprog (SUC foo) ((h,z),r) = x)}) UNION
3805                 (IMAGE (\(h,r). ((h, z),r))
3806                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3807                                ~r (STRCAT "coin" (toString 1)) /\
3808                                (insider_dcprog (SUC foo) ((h,z),r) = x)})`
3809                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
3810                                    IN_CROSS, L_def, IN_UNION]
3811                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
3812                     >- (Cases_on `SND x' (STRCAT "coin" (toString 1))`
3813                         >- (DISJ1_TAC >> Q.EXISTS_TAC `(FST(FST x'), SND x')`
3814                             >> RW_TAC std_ss [PAIR]
3815                             >> Q.EXISTS_TAC `(FST (FST x'),SND x')`
3816                             >> RW_TAC std_ss [PAIR_EQ])
3817                         >> DISJ2_TAC >> Q.EXISTS_TAC `(FST(FST x'), SND x')`
3818                         >> RW_TAC std_ss [PAIR]
3819                         >> Q.EXISTS_TAC `(FST (FST x'),SND x')`
3820                         >> RW_TAC std_ss [PAIR_EQ])
3821                     >>  POP_ASSUM MP_TAC
3822                     >> `x''' = (FST x''',SND x''')` by RW_TAC std_ss [PAIR]
3823                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
3824                     >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3825                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
3826            >> POP_ORW
3827            >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3828                                r (STRCAT "coin" (toString 1)) /\
3829                                (insider_dcprog (SUC foo) ((h,z),r) = x)} /\
3830                FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3831                                ~r (STRCAT "coin" (toString 1)) /\
3832                                (insider_dcprog (SUC foo) ((h,z),r) = x)}`
3833                by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
3834                     >> RW_TAC bool_ss [FINITE_CROSS]
3835                     >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
3836                     >> POP_ASSUM MP_TAC
3837                     >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3838                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
3839                     >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3840                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
3841            >> `DISJOINT (IMAGE (\(h,r). ((h, z),r))
3842                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3843                                r (STRCAT "coin" (toString 1)) /\
3844                                (insider_dcprog (SUC foo) ((h,z),r) = x)})
3845                         (IMAGE (\(h,r). ((h, z),r))
3846                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3847                                ~r (STRCAT "coin" (toString 1)) /\
3848                                (insider_dcprog (SUC foo) ((h,z),r) = x)})`
3849                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, IN_IMAGE, GSPECIFICATION]
3850                    >> Cases_on `SND x' (STRCAT "coin" (toString 1))`
3851                    >- (DISJ2_TAC >> STRIP_TAC >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3852                        >> POP_ORW >> RW_TAC std_ss [] >> REVERSE (Cases_on `x' = ((FST x'',z),SND x'')`) >> RW_TAC std_ss []
3853                        >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3854                        >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3855                        >> METIS_TAC [FST, SND, PAIR_EQ, PAIR])
3856                    >> DISJ1_TAC >> STRIP_TAC >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3857                    >> POP_ORW >> RW_TAC std_ss [] >> REVERSE (Cases_on `x' = ((FST x'',z),SND x'')`) >> RW_TAC std_ss []
3858                    >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3859                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3860                    >> METIS_TAC [FST, SND, PAIR_EQ, PAIR])
3861            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE] >> POP_ASSUM (K ALL_TAC)
3862            >> `INJ (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3863                                r (STRCAT "coin" (toString 1)) /\
3864                                (insider_dcprog (SUC foo) ((h,z),r) = x)}
3865                    (IMAGE (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3866                                r (STRCAT "coin" (toString 1)) /\
3867                                (insider_dcprog (SUC foo) ((h,z),r) = x)}) /\
3868                INJ (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3869                                ~r (STRCAT "coin" (toString 1)) /\
3870                                (insider_dcprog (SUC foo) ((h,z),r) = x)}
3871                    (IMAGE (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3872                                ~r (STRCAT "coin" (toString 1)) /\
3873                                (insider_dcprog (SUC foo) ((h,z),r) = x)})`
3874                by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE]
3875                    >| [METIS_TAC [],
3876                        POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
3877                            >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3878                            >> POP_ORW >> RW_TAC std_ss []
3879                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
3880                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
3881                            >> POP_ORW >> RW_TAC std_ss []
3882                            >> METIS_TAC [PAIR, PAIR_EQ],
3883                        METIS_TAC [],
3884                        POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
3885                            >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3886                            >> POP_ORW >> RW_TAC std_ss []
3887                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
3888                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
3889                            >> POP_ORW >> RW_TAC std_ss []
3890                            >> METIS_TAC [PAIR, PAIR_EQ]])
3891            >> RW_TAC bool_ss [REAL_SUM_IMAGE_IMAGE, IMAGE_FINITE, o_DEF] >> POP_ASSUM (K ALL_TAC) >> POP_ASSUM (K ALL_TAC)
3892            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
3893            >> `!x': bool state # bool state.
3894                (if x' IN {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3895                                   r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}
3896                then (if SND ((\(h,r). ((h,z),r)) x') (STRCAT "coin" (toString 1)) then
3897                1 / 2 * (if (\(h,r). ((h,z),r)) x' IN biased_high_states foo CROSS biased_low_states CROSS
3898                                                        biased_random_states foo
3899                then 1 / 16 else 0) else 3 / 2 *
3900                (if (\(h,r). ((h,z),r)) x' IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
3901                then 1 / 16 else 0)) else 0) =
3902                1/32 * (\x'. if
3903                 x' IN {(h,r) |
3904                    h IN biased_high_states foo /\ r IN biased_random_states foo /\
3905                       r (STRCAT "coin" (toString 1)) /\
3906                (insider_dcprog (SUC foo) ((h,z),r) = x)}
3907                then (\x'. 1) x' else 0) x'`
3908                by (STRIP_TAC >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3909                    >> POP_ORW >> RW_TAC real_ss [GSPECIFICATION, IN_CROSS]
3910                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss []
3911                    >> `(x'':bool state # bool state) = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3912                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3913                    >> METIS_TAC [])
3914            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
3915            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
3916            >> `!x': bool state # bool state.
3917                (if x' IN {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3918                                   ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}
3919                then (if SND ((\(h,r). ((h,z),r)) x') (STRCAT "coin" (toString 1)) then
3920                1 / 2 * (if (\(h,r). ((h,z),r)) x' IN biased_high_states foo CROSS biased_low_states CROSS
3921                                                        biased_random_states foo
3922                then 1 / 16 else 0) else 3 / 2 *
3923                (if (\(h,r). ((h,z),r)) x' IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
3924                then 1 / 16 else 0)) else 0) =
3925                3/32 * (\x'. if
3926                 x' IN {(h,r) |
3927                    h IN biased_high_states foo /\ r IN biased_random_states foo /\
3928                       ~r (STRCAT "coin" (toString 1)) /\
3929                (insider_dcprog (SUC foo) ((h,z),r) = x)}
3930                then (\x'. 1) x' else 0) x'`
3931                by (STRIP_TAC >> `(x' :bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
3932                    >> POP_ORW >> RW_TAC real_ss [GSPECIFICATION, IN_CROSS]
3933                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss []
3934                    >> `(x'':bool state # bool state) = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
3935                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
3936                    >> METIS_TAC [])
3937            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
3938            >> RW_TAC real_ss [])
3939   >> `!x z. (PREIMAGE (\x. (insider_dcprog (SUC foo) x,SND (FST x)))
3940           {(x,z)} INTER
3941         (biased_high_states foo CROSS biased_low_states CROSS
3942          biased_random_states foo)) =
3943        (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
3944           PREIMAGE L {z} INTER
3945           (biased_high_states foo CROSS biased_low_states CROSS
3946            biased_random_states foo))`
3947        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, IN_CROSS, IN_PREIMAGE, IN_SING, L_def])
3948   >> POP_ORW
3949   >> POP_ORW
3950   >> `!x z.
3951        SIGMA
3952          (\s.
3953             (if SND s (STRCAT "coin" (toString 1)) then
3954                1 / 2 *
3955                (if
3956                   s IN
3957                   biased_high_states foo CROSS biased_low_states CROSS
3958                   biased_random_states foo
3959                 then
3960                   1 / 16
3961                 else
3962                   0)
3963              else
3964                3 / 2 *
3965                (if
3966                   s IN
3967                   biased_high_states foo CROSS biased_low_states CROSS
3968                   biased_random_states foo
3969                 then
3970                   1 / 16
3971                 else
3972                   0)))
3973          (PREIMAGE L {z} INTER
3974           (biased_high_states foo CROSS biased_low_states CROSS
3975            biased_random_states foo)) =
3976        (if z IN biased_low_states then
3977                1/32 * & (CARD (biased_high_states foo)) *
3978                        &(CARD {r| r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}) +
3979                3/32 * & (CARD (biased_high_states foo)) *
3980                        &(CARD {r| r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))})
3981         else 0)`
3982        by (REVERSE (RW_TAC real_ss [])
3983            >- (Suff `PREIMAGE L {z} INTER (biased_high_states foo CROSS biased_low_states CROSS
3984                                                biased_random_states foo) = {}`
3985                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
3986                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, L_def, IN_CROSS, IN_INTER, IN_PREIMAGE]
3987                >> METIS_TAC [])
3988            >> `(PREIMAGE L {z} INTER
3989                (biased_high_states foo CROSS biased_low_states CROSS
3990                 biased_random_states foo)) =
3991                 (IMAGE (\(h,r). ((h, z),r))
3992                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3993                                r (STRCAT "coin" (toString 1))}) UNION
3994                 (IMAGE (\(h,r). ((h, z),r))
3995                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
3996                                ~r (STRCAT "coin" (toString 1))})`
3997                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
3998                                    IN_CROSS, L_def, IN_UNION]
3999                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4000                     >- (Cases_on `SND x (STRCAT "coin" (toString 1))`
4001                         >- (DISJ1_TAC >> Q.EXISTS_TAC `(FST(FST x), SND x)`
4002                             >> RW_TAC std_ss [PAIR]
4003                             >> Q.EXISTS_TAC `(FST (FST x),SND x)`
4004                             >> RW_TAC std_ss [PAIR_EQ])
4005                         >> DISJ2_TAC >> Q.EXISTS_TAC `(FST(FST x), SND x)`
4006                         >> RW_TAC std_ss [PAIR]
4007                         >> Q.EXISTS_TAC `(FST (FST x),SND x)`
4008                         >> RW_TAC std_ss [PAIR_EQ])
4009                     >>  POP_ASSUM MP_TAC
4010                     >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
4011                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
4012                     >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4013                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
4014            >> POP_ORW
4015            >> `FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4016                                r (STRCAT "coin" (toString 1))} /\
4017                FINITE {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4018                                ~r (STRCAT "coin" (toString 1))}`
4019                by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
4020                     >> RW_TAC bool_ss [FINITE_CROSS]
4021                     >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
4022                     >> POP_ASSUM MP_TAC
4023                     >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4024                     >> POP_ORW >> RW_TAC std_ss [FST,SND, PAIR_EQ]
4025                     >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4026                     >> POP_ORW >> RW_TAC std_ss [] >> RW_TAC std_ss [FST,SND])
4027            >> `DISJOINT (IMAGE (\(h,r). ((h, z),r))
4028                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4029                                r (STRCAT "coin" (toString 1))})
4030                         (IMAGE (\(h,r). ((h, z),r))
4031                        {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4032                                ~r (STRCAT "coin" (toString 1))})`
4033                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, IN_IMAGE, GSPECIFICATION]
4034                    >> Cases_on `SND x (STRCAT "coin" (toString 1))`
4035                    >- (DISJ2_TAC >> STRIP_TAC >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4036                        >> POP_ORW >> RW_TAC std_ss [] >> REVERSE (Cases_on `x = ((FST x',z),SND x')`) >> RW_TAC std_ss []
4037                        >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4038                        >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4039                        >> METIS_TAC [FST, SND, PAIR_EQ, PAIR])
4040                    >> DISJ1_TAC >> STRIP_TAC >> `x' = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4041                    >> POP_ORW >> RW_TAC std_ss [] >> REVERSE (Cases_on `x = ((FST x',z),SND x')`) >> RW_TAC std_ss []
4042                    >> `x = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4043                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4044                    >> METIS_TAC [FST, SND, PAIR_EQ, PAIR])
4045            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE] >> POP_ASSUM (K ALL_TAC)
4046            >> `INJ (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4047                                r (STRCAT "coin" (toString 1))}
4048                    (IMAGE (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4049                                r (STRCAT "coin" (toString 1))}) /\
4050                INJ (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4051                                ~r (STRCAT "coin" (toString 1))}
4052                    (IMAGE (\(h,r). ((h,z),r)) {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4053                                ~r (STRCAT "coin" (toString 1))})`
4054                by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE]
4055                    >| [METIS_TAC [],
4056                        POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
4057                            >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4058                            >> POP_ORW >> RW_TAC std_ss []
4059                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
4060                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
4061                            >> POP_ORW >> RW_TAC std_ss []
4062                            >> METIS_TAC [PAIR, PAIR_EQ],
4063                        METIS_TAC [],
4064                        POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
4065                            >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4066                            >> POP_ORW >> RW_TAC std_ss []
4067                            >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC >> POP_ASSUM MP_TAC
4068                            >> `(y :bool state # bool state) = (FST y,SND y)` by RW_TAC std_ss [PAIR]
4069                            >> POP_ORW >> RW_TAC std_ss []
4070                            >> METIS_TAC [PAIR, PAIR_EQ]])
4071            >> RW_TAC bool_ss [REAL_SUM_IMAGE_IMAGE, IMAGE_FINITE, o_DEF] >> POP_ASSUM (K ALL_TAC) >> POP_ASSUM (K ALL_TAC)
4072            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4073            >> `!x: bool state # bool state.
4074                (if x IN {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4075                                   r (STRCAT "coin" (toString 1))}
4076                then (if SND ((\(h,r). ((h,z),r)) x) (STRCAT "coin" (toString 1)) then
4077                1 / 2 * (if (\(h,r). ((h,z),r)) x IN biased_high_states foo CROSS biased_low_states CROSS
4078                                                        biased_random_states foo
4079                then 1 / 16 else 0) else 3 / 2 *
4080                (if (\(h,r). ((h,z),r)) x IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4081                then 1 / 16 else 0)) else 0) =
4082                1/32 * (\x. if
4083                 x IN {(h,r) |
4084                    h IN biased_high_states foo /\ r IN biased_random_states foo /\
4085                       r (STRCAT "coin" (toString 1))}
4086                then (\x. 1) x else 0) x`
4087                by (STRIP_TAC >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4088                    >> POP_ORW >> RW_TAC real_ss [GSPECIFICATION, IN_CROSS]
4089                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss []
4090                    >> `(x':bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4091                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4092                    >> METIS_TAC [])
4093            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
4094            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4095            >> `!x: bool state # bool state.
4096                (if x IN {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4097                                   ~r (STRCAT "coin" (toString 1))}
4098                then (if SND ((\(h,r). ((h,z),r)) x) (STRCAT "coin" (toString 1)) then
4099                1 / 2 * (if (\(h,r). ((h,z),r)) x IN biased_high_states foo CROSS biased_low_states CROSS
4100                                                        biased_random_states foo
4101                then 1 / 16 else 0) else 3 / 2 *
4102                (if (\(h,r). ((h,z),r)) x IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4103                then 1 / 16 else 0)) else 0) =
4104                3/32 * (\x'. if
4105                 x IN {(h,r) |
4106                    h IN biased_high_states foo /\ r IN biased_random_states foo /\
4107                       ~r (STRCAT "coin" (toString 1))}
4108                then (\x. 1) x else 0) x`
4109                by (STRIP_TAC >> `(x :bool state # bool state) = (FST x,SND x)` by RW_TAC std_ss [PAIR]
4110                    >> POP_ORW >> RW_TAC real_ss [GSPECIFICATION, IN_CROSS]
4111                    >> POP_ASSUM MP_TAC >> RW_TAC std_ss []
4112                    >> `(x':bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4113                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4114                    >> METIS_TAC [])
4115            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
4116            >> RW_TAC real_ss []
4117            >> `{(h,r) |  h IN biased_high_states foo /\ r IN biased_random_states foo /\
4118                          r (STRCAT "coin" (toString 1))} =
4119                (biased_high_states foo) CROSS {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}`
4120                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [GSPECIFICATION, IN_CROSS]
4121                    >> REVERSE (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4122                    >- (Q.EXISTS_TAC `(FST x, SND x)` >> RW_TAC std_ss [PAIR_EQ])
4123                    >> POP_ASSUM MP_TAC >> `(x':bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4124                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ])
4125            >> POP_ORW
4126            >> `{(h,r) |  h IN biased_high_states foo /\ r IN biased_random_states foo /\
4127                          ~r (STRCAT "coin" (toString 1))} =
4128                (biased_high_states foo) CROSS {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4129                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [GSPECIFICATION, IN_CROSS]
4130                    >> REVERSE (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4131                    >- (Q.EXISTS_TAC `(FST x, SND x)` >> RW_TAC std_ss [PAIR_EQ])
4132                    >> POP_ASSUM MP_TAC >> `(x':bool state # bool state) = (FST x',SND x')` by RW_TAC std_ss [PAIR]
4133                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ])
4134             >> POP_ORW
4135             >> `FINITE {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))} /\
4136                 FINITE {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4137                by (CONJ_TAC
4138                    >> (MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4139                    >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4140             >> RW_TAC std_ss [CARD_CROSS, GSYM REAL_MUL, REAL_MUL_ASSOC, REAL_MUL_COMM])
4141   >> POP_ORW
4142   >> `!x y z.
4143        SIGMA
4144          (\s.
4145             (if SND s (STRCAT "coin" (toString 1)) then
4146                1 / 2 *
4147                (if
4148                   s IN
4149                   biased_high_states foo CROSS biased_low_states CROSS
4150                   biased_random_states foo
4151                 then
4152                   1 / 16
4153                 else
4154                   0)
4155              else
4156                3 / 2 *
4157                (if
4158                   s IN
4159                   biased_high_states foo CROSS biased_low_states CROSS
4160                   biased_random_states foo
4161                 then
4162                   1 / 16
4163                 else
4164                   0)))
4165          (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
4166           PREIMAGE FST {(y,z)} INTER
4167           (biased_high_states foo CROSS biased_low_states CROSS
4168            biased_random_states foo)) =
4169        (if z IN biased_low_states /\ y IN (biased_high_states foo) then
4170                1/32 * & (CARD {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1)) /\
4171                                    (insider_dcprog (SUC foo) ((y,z),r) = x)})
4172                + 3/32 * & (CARD {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1)) /\
4173                                      (insider_dcprog (SUC foo) ((y,z),r) = x)})
4174         else 0)`
4175        by (REVERSE (RW_TAC real_ss [])
4176            >- (RW_TAC bool_ss [GSYM INTER_ASSOC]
4177                >> Suff `PREIMAGE FST {(y,z)} INTER (biased_high_states foo CROSS biased_low_states CROSS
4178                                                biased_random_states foo) = {}`
4179                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
4180                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, FST, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING]
4181                >> METIS_TAC [PAIR, FST,SND])
4182            >> `(PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
4183                 PREIMAGE FST {(y,z)} INTER
4184                (biased_high_states foo CROSS biased_low_states CROSS
4185                 biased_random_states foo)) =
4186                 (IMAGE (\r. ((y, z),r))
4187                        {r | r IN biased_random_states foo /\
4188                                r (STRCAT "coin" (toString 1)) /\
4189                                (insider_dcprog (SUC foo) ((y,z),r) = x)}) UNION
4190                 (IMAGE (\r. ((y, z),r))
4191                        {r | r IN biased_random_states foo /\
4192                                ~r (STRCAT "coin" (toString 1)) /\
4193                                (insider_dcprog (SUC foo) ((y,z),r) = x)})`
4194                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
4195                                    IN_CROSS, FST, IN_UNION]
4196                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4197                     >> Cases_on `SND x' (STRCAT "coin" (toString 1))`
4198                     >- (DISJ1_TAC >> Q.EXISTS_TAC `SND x'`
4199                         >> METIS_TAC [PAIR])
4200                     >> DISJ2_TAC >> Q.EXISTS_TAC `SND x'`
4201                     >> METIS_TAC [PAIR])
4202            >> POP_ORW
4203            >> `FINITE {r | r IN biased_random_states foo /\
4204                                r (STRCAT "coin" (toString 1)) /\
4205                                (insider_dcprog (SUC foo) ((y,z),r) = x)} /\
4206                FINITE {r | r IN biased_random_states foo /\
4207                                ~r (STRCAT "coin" (toString 1)) /\
4208                                (insider_dcprog (SUC foo) ((y,z),r) = x)}`
4209                by (CONJ_TAC
4210                    >> (MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4211                    >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4212            >> `DISJOINT (IMAGE (\r. ((y, z),r))
4213                        {r | r IN biased_random_states foo /\
4214                                r (STRCAT "coin" (toString 1)) /\
4215                                (insider_dcprog (SUC foo) ((y,z),r) = x)})
4216                         (IMAGE (\r. ((y, z),r))
4217                        {r | r IN biased_random_states foo /\
4218                                ~r (STRCAT "coin" (toString 1)) /\
4219                                (insider_dcprog (SUC foo) ((y,z),r) = x)})`
4220                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, IN_IMAGE, GSPECIFICATION]
4221                    >> METIS_TAC [PAIR_EQ])
4222            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE] >> POP_ASSUM (K ALL_TAC)
4223            >> `INJ (\r. ((y, z),r))
4224                        {r | r IN biased_random_states foo /\
4225                                r (STRCAT "coin" (toString 1)) /\
4226                                (insider_dcprog (SUC foo) ((y,z),r) = x)}
4227                    (IMAGE (\r. ((y, z),r))
4228                        {r | r IN biased_random_states foo /\
4229                                r (STRCAT "coin" (toString 1)) /\
4230                                (insider_dcprog (SUC foo) ((y,z),r) = x)}) /\
4231                INJ (\r. ((y, z),r))
4232                        {r | r IN biased_random_states foo /\
4233                                ~r (STRCAT "coin" (toString 1)) /\
4234                                (insider_dcprog (SUC foo) ((y,z),r) = x)}
4235                    (IMAGE (\r. ((y, z),r))
4236                        {r | r IN biased_random_states foo /\
4237                                ~r (STRCAT "coin" (toString 1)) /\
4238                                (insider_dcprog (SUC foo) ((y,z),r) = x)})`
4239                by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
4240            >> RW_TAC bool_ss [REAL_SUM_IMAGE_IMAGE, IMAGE_FINITE, o_DEF] >> POP_ASSUM (K ALL_TAC) >> POP_ASSUM (K ALL_TAC)
4241            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4242            >> `!r: bool state.
4243                (if r IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1)) /\
4244                              (insider_dcprog (SUC foo) ((y,z),r) = x)}
4245                then (if SND ((y,z),r) (STRCAT "coin" (toString 1)) then
4246                1 / 2 * (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS
4247                                                        biased_random_states foo
4248                then 1 / 16 else 0) else 3 / 2 *
4249                (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4250                then 1 / 16 else 0)) else 0) =
4251                1/32 * (\r. if
4252                 r IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1)) /\
4253                              (insider_dcprog (SUC foo) ((y,z),r) = x)}
4254                then (\r. 1) r else 0) r`
4255                by (RW_TAC real_ss [GSPECIFICATION, IN_CROSS])
4256            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
4257            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4258            >> `!r: bool state.
4259                (if r IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1)) /\
4260                              (insider_dcprog (SUC foo) ((y,z),r) = x)}
4261                then (if r (STRCAT "coin" (toString 1)) then
4262                1 / 2 * (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS
4263                                                        biased_random_states foo
4264                then 1 / 16 else 0) else 3 / 2 *
4265                (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4266                then 1 / 16 else 0)) else 0) =
4267                3/32 * (\r. if
4268                 r IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1)) /\
4269                              (insider_dcprog (SUC foo) ((y,z),r) = x)}
4270                then (\r. 1) r else 0) r`
4271                by (RW_TAC real_ss [GSPECIFICATION, IN_CROSS])
4272            >> POP_ORW >> RW_TAC real_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3])
4273   >> `!x y z. (PREIMAGE (\x. (insider_dcprog (SUC foo) x,FST x))
4274           {(x,y,z)} INTER
4275         (biased_high_states foo CROSS biased_low_states CROSS
4276          biased_random_states foo)) =
4277        (PREIMAGE (insider_dcprog (SUC foo)) {x} INTER
4278           PREIMAGE FST {(y,z)} INTER
4279           (biased_high_states foo CROSS biased_low_states CROSS
4280            biased_random_states foo))`
4281        by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, IN_CROSS, IN_PREIMAGE, IN_SING, L_def])
4282   >> POP_ORW
4283   >> POP_ORW
4284   >> `!x y z.
4285        SIGMA
4286          (\s.
4287             (if SND s (STRCAT "coin" (toString 1)) then
4288                1 / 2 *
4289                (if
4290                   s IN
4291                   biased_high_states foo CROSS biased_low_states CROSS
4292                   biased_random_states foo
4293                 then
4294                   1 / 16
4295                 else
4296                   0)
4297              else
4298                3 / 2 *
4299                (if
4300                   s IN
4301                   biased_high_states foo CROSS biased_low_states CROSS
4302                   biased_random_states foo
4303                 then
4304                   1 / 16
4305                 else
4306                   0)))
4307          (PREIMAGE FST {(y,z)} INTER
4308           (biased_high_states foo CROSS biased_low_states CROSS
4309            biased_random_states foo)) =
4310        (if z IN biased_low_states /\ y IN (biased_high_states foo) then
4311                1/32 * & (CARD {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))})
4312                + 3/32 * & (CARD {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))})
4313         else 0)`
4314        by (REVERSE (RW_TAC real_ss [])
4315            >- (Suff `PREIMAGE FST {(y,z)} INTER (biased_high_states foo CROSS biased_low_states CROSS
4316                                                biased_random_states foo) = {}`
4317                >- RW_TAC bool_ss [INTER_EMPTY, REAL_SUM_IMAGE_THM]
4318                >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC set_ss [NOT_IN_EMPTY, FST, IN_CROSS, IN_INTER, IN_PREIMAGE, IN_SING]
4319                >> METIS_TAC [PAIR, FST,SND])
4320            >> `(PREIMAGE FST {(y,z)} INTER
4321                (biased_high_states foo CROSS biased_low_states CROSS
4322                 biased_random_states foo)) =
4323                 (IMAGE (\r. ((y, z),r))
4324                        {r | r IN biased_random_states foo /\
4325                                r (STRCAT "coin" (toString 1))}) UNION
4326                 (IMAGE (\r. ((y, z),r))
4327                        {r | r IN biased_random_states foo /\
4328                                ~r (STRCAT "coin" (toString 1))})`
4329                 by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, GSPECIFICATION, IN_IMAGE, IN_PREIMAGE, IN_SING,
4330                                    IN_CROSS, FST, IN_UNION]
4331                     >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4332                     >> METIS_TAC [PAIR])
4333            >> POP_ORW
4334            >> `FINITE {r | r IN biased_random_states foo /\
4335                                r (STRCAT "coin" (toString 1))} /\
4336                FINITE {r | r IN biased_random_states foo /\
4337                                ~r (STRCAT "coin" (toString 1))}`
4338                by (CONJ_TAC
4339                    >> (MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4340                    >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4341            >> `DISJOINT (IMAGE (\r. ((y, z),r))
4342                        {r | r IN biased_random_states foo /\
4343                                r (STRCAT "coin" (toString 1))})
4344                         (IMAGE (\r. ((y, z),r))
4345                        {r | r IN biased_random_states foo /\
4346                                ~r (STRCAT "coin" (toString 1))})`
4347                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, IN_IMAGE, GSPECIFICATION]
4348                    >> METIS_TAC [PAIR_EQ])
4349            >> RW_TAC bool_ss [REAL_SUM_IMAGE_DISJOINT_UNION, IMAGE_FINITE] >> POP_ASSUM (K ALL_TAC)
4350            >> `INJ (\r. ((y, z),r))
4351                        {r | r IN biased_random_states foo /\
4352                                r (STRCAT "coin" (toString 1))}
4353                    (IMAGE (\r. ((y, z),r))
4354                        {r | r IN biased_random_states foo /\
4355                                r (STRCAT "coin" (toString 1))}) /\
4356                INJ (\r. ((y, z),r))
4357                        {r | r IN biased_random_states foo /\
4358                                ~r (STRCAT "coin" (toString 1))}
4359                    (IMAGE (\r. ((y, z),r))
4360                        {r | r IN biased_random_states foo /\
4361                                ~r (STRCAT "coin" (toString 1))})`
4362                by (RW_TAC std_ss [INJ_DEF, GSPECIFICATION, IN_IMAGE])
4363            >> RW_TAC bool_ss [REAL_SUM_IMAGE_IMAGE, IMAGE_FINITE, o_DEF] >> POP_ASSUM (K ALL_TAC) >> POP_ASSUM (K ALL_TAC)
4364            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4365            >> `!r: bool state.
4366                (if r IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4367                then (if SND ((y,z),r) (STRCAT "coin" (toString 1)) then
4368                1 / 2 * (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS
4369                                                        biased_random_states foo
4370                then 1 / 16 else 0) else 3 / 2 *
4371                (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4372                then 1 / 16 else 0)) else 0) =
4373                1/32 * (\r. if
4374                 r IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4375                then (\r. 1) r else 0) r`
4376                by (RW_TAC real_ss [GSPECIFICATION, IN_CROSS])
4377            >> POP_ORW >> RW_TAC bool_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3]
4378            >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF]
4379            >> `!r: bool state.
4380                (if r IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}
4381                then (if r (STRCAT "coin" (toString 1)) then
4382                1 / 2 * (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS
4383                                                        biased_random_states foo
4384                then 1 / 16 else 0) else 3 / 2 *
4385                (if ((y,z),r) IN biased_high_states foo CROSS biased_low_states CROSS biased_random_states foo
4386                then 1 / 16 else 0)) else 0) =
4387                3/32 * (\r. if
4388                 r IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}
4389                then (\r. 1) r else 0) r`
4390                by (RW_TAC real_ss [GSPECIFICATION, IN_CROSS])
4391            >> POP_ORW >> RW_TAC real_ss [REAL_SUM_IMAGE_CMUL, GSYM REAL_SUM_IMAGE_IN_IF, REAL_SUM_IMAGE_FINITE_CONST3])
4392   >> POP_ORW
4393   >> `FINITE (IMAGE (insider_dcprog (SUC foo))
4394             (biased_high_states foo CROSS biased_low_states CROSS
4395              biased_random_states foo) CROSS biased_low_states)`
4396        by RW_TAC std_ss [IMAGE_FINITE, FINITE_CROSS]
4397   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `IMAGE (insider_dcprog (SUC foo))
4398             (biased_high_states foo CROSS biased_low_states CROSS
4399              biased_random_states foo) CROSS biased_low_states`) REAL_SUM_IMAGE_IN_IF]
4400   >> Q.UNABBREV_TAC `foo`
4401   >> RW_TAC bool_ss [(COMPUTE_CARD ``(biased_high_states (SUC (SUC 0)))``
4402                              (SIMP_CONV set_ss [biased_high_states_def, insider_high_states_set_def]
4403                               THENC (TRY_CONV insider_hl_dup_conv))
4404                                dc_dup_conv)]
4405   >> Q.ABBREV_TAC `foo = SUC (SUC 0)`
4406   >> RW_TAC real_ss []
4407   >> RW_TAC bool_ss [Once REAL_SUM_IMAGE_IN_IF, IMAGE_FINITE, FINITE_CROSS]
4408   >> `!x:bool state # bool state.
4409        (if
4410           x IN
4411           IMAGE (insider_dcprog (SUC foo))
4412             (biased_high_states foo CROSS biased_low_states CROSS
4413              biased_random_states foo) CROSS biased_low_states
4414         then
4415           (\(x,z).
4416              (if z IN biased_low_states then
4417                 1 / 32 *
4418                 &
4419                   (CARD
4420                      {(h,r) |
4421                       h IN biased_high_states foo /\
4422                       r IN biased_random_states foo /\
4423                       r (STRCAT "coin" (toString 1)) /\
4424                       (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4425                 3 / 32 *
4426                 &
4427                   (CARD
4428                      {(h,r) |
4429                       h IN biased_high_states foo /\
4430                       r IN biased_random_states foo /\
4431                       ~r (STRCAT "coin" (toString 1)) /\
4432                       (insider_dcprog (SUC foo) ((h,z),r) = x)})
4433               else
4434                 0) *
4435              logr 2
4436                ((if z IN biased_low_states then
4437                    1 / 32 *
4438                    &
4439                      (CARD
4440                         {(h,r) |
4441                          h IN biased_high_states foo /\
4442                          r IN biased_random_states foo /\
4443                          r (STRCAT "coin" (toString 1)) /\
4444                          (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4445                    3 / 32 *
4446                    &
4447                      (CARD
4448                         {(h,r) |
4449                          h IN biased_high_states foo /\
4450                          r IN biased_random_states foo /\
4451                          ~r (STRCAT "coin" (toString 1)) /\
4452                          (insider_dcprog (SUC foo) ((h,z),r) = x)})
4453                  else
4454                    0) /
4455                 (if z IN biased_low_states then
4456                    1 / 16 *
4457                    &
4458                      (CARD
4459                         {r |
4460                          r IN biased_random_states foo /\
4461                          r (STRCAT "coin" (toString 1))}) +
4462                    3 / 16 *
4463                    &
4464                      (CARD
4465                         {r |
4466                          r IN biased_random_states foo /\
4467                          ~r (STRCAT "coin" (toString 1))})
4468                  else
4469                    0))) x
4470         else
4471           0) =
4472        (if
4473           x IN
4474           IMAGE (insider_dcprog (SUC foo))
4475             (biased_high_states foo CROSS biased_low_states CROSS
4476              biased_random_states foo) CROSS biased_low_states
4477         then
4478           (\(x,z).
4479              (1 / 32 *
4480                 &
4481                   (CARD
4482                      {(h,r) |
4483                       h IN biased_high_states foo /\
4484                       r IN biased_random_states foo /\
4485                       r (STRCAT "coin" (toString 1)) /\
4486                       (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4487                 3 / 32 *
4488                 &
4489                   (CARD
4490                      {(h,r) |
4491                       h IN biased_high_states foo /\
4492                       r IN biased_random_states foo /\
4493                       ~r (STRCAT "coin" (toString 1)) /\
4494                       (insider_dcprog (SUC foo) ((h,z),r) = x)})) *
4495              logr 2
4496                ((1 / 32 *
4497                    &
4498                      (CARD
4499                         {(h,r) |
4500                          h IN biased_high_states foo /\
4501                          r IN biased_random_states foo /\
4502                          r (STRCAT "coin" (toString 1)) /\
4503                          (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4504                    3 / 32 *
4505                    &
4506                      (CARD
4507                         {(h,r) |
4508                          h IN biased_high_states foo /\
4509                          r IN biased_random_states foo /\
4510                          ~r (STRCAT "coin" (toString 1)) /\
4511                          (insider_dcprog (SUC foo) ((h,z),r) = x)})) /
4512                 (1 / 16 *
4513                    &
4514                      (CARD
4515                         {r |
4516                          r IN biased_random_states foo /\
4517                          r (STRCAT "coin" (toString 1))}) +
4518                    3 / 16 *
4519                    &
4520                      (CARD
4521                         {r |
4522                          r IN biased_random_states foo /\
4523                          ~r (STRCAT "coin" (toString 1))})))) x
4524         else
4525           0)`
4526        by (RW_TAC std_ss [IN_CROSS] >> `x = (FST x, SND x)` by RW_TAC std_ss [PAIR] >> POP_ORW >> RW_TAC std_ss [])
4527   >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF, IMAGE_FINITE, FINITE_CROSS]
4528   >> Q.ABBREV_TAC `sumone =
4529        SIGMA
4530     (\x.
4531        (\(x,z).
4532           (1 / 32 *
4533            &
4534              (CARD
4535                 {(h,r) |
4536                  h IN biased_high_states foo /\
4537                  r IN biased_random_states foo /\
4538                  r (STRCAT "coin" (toString 1)) /\
4539                  (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4540            3 / 32 *
4541            &
4542              (CARD
4543                 {(h,r) |
4544                  h IN biased_high_states foo /\
4545                  r IN biased_random_states foo /\
4546                  ~r (STRCAT "coin" (toString 1)) /\
4547                  (insider_dcprog (SUC foo) ((h,z),r) = x)})) *
4548           logr 2
4549             ((1 / 32 *
4550               &
4551                 (CARD
4552                    {(h,r) |
4553                     h IN biased_high_states foo /\
4554                     r IN biased_random_states foo /\
4555                     r (STRCAT "coin" (toString 1)) /\
4556                     (insider_dcprog (SUC foo) ((h,z),r) = x)}) +
4557               3 / 32 *
4558               &
4559                 (CARD
4560                    {(h,r) |
4561                     h IN biased_high_states foo /\
4562                     r IN biased_random_states foo /\
4563                     ~r (STRCAT "coin" (toString 1)) /\
4564                     (insider_dcprog (SUC foo) ((h,z),r) = x)})) /
4565              (1 / 16 *
4566               &
4567                 (CARD
4568                    {r |
4569                     r IN biased_random_states foo /\
4570                     r (STRCAT "coin" (toString 1))}) +
4571               3 / 16 *
4572               &
4573                 (CARD
4574                    {r |
4575                     r IN biased_random_states foo /\
4576                     ~r (STRCAT "coin" (toString 1))})))) x)
4577     (IMAGE (insider_dcprog (SUC foo))
4578        (biased_high_states foo CROSS biased_low_states CROSS
4579         biased_random_states foo) CROSS biased_low_states)`
4580   >> `FINITE (IMAGE (insider_dcprog (SUC foo))
4581            (biased_high_states foo CROSS biased_low_states CROSS
4582             biased_random_states foo) CROSS
4583          (biased_high_states foo CROSS biased_low_states))`
4584        by RW_TAC std_ss [IMAGE_FINITE, FINITE_CROSS]
4585   >> ONCE_REWRITE_TAC [(UNDISCH o Q.ISPEC `IMAGE (insider_dcprog (SUC foo))
4586            (biased_high_states foo CROSS biased_low_states CROSS
4587             biased_random_states foo) CROSS
4588          (biased_high_states foo CROSS biased_low_states)`) REAL_SUM_IMAGE_IN_IF]
4589   >> `!x. (if
4590          x IN
4591          IMAGE (insider_dcprog (SUC foo))
4592            (biased_high_states foo CROSS biased_low_states CROSS
4593             biased_random_states foo) CROSS
4594          (biased_high_states foo CROSS biased_low_states)
4595        then
4596          (\(x,y,z).
4597             (if
4598                z IN biased_low_states /\ y IN biased_high_states foo
4599              then
4600                1 / 32 *
4601                &
4602                  (CARD
4603                     {r |
4604                      r IN biased_random_states foo /\
4605                      r (STRCAT "coin" (toString 1)) /\
4606                      (insider_dcprog (SUC foo) ((y,z),r) = x)}) +
4607                3 / 32 *
4608                &
4609                  (CARD
4610                     {r |
4611                      r IN biased_random_states foo /\
4612                      ~r (STRCAT "coin" (toString 1)) /\
4613                      (insider_dcprog (SUC foo) ((y,z),r) = x)})
4614              else
4615                0) *
4616             logr 2
4617               ((if
4618                   z IN biased_low_states /\ y IN biased_high_states foo
4619                 then
4620                   1 / 32 *
4621                   &
4622                     (CARD
4623                        {r |
4624                         r IN biased_random_states foo /\
4625                         r (STRCAT "coin" (toString 1)) /\
4626                         (insider_dcprog (SUC foo) ((y,z),r) = x)}) +
4627                   3 / 32 *
4628                   &
4629                     (CARD
4630                        {r |
4631                         r IN biased_random_states foo /\
4632                         ~r (STRCAT "coin" (toString 1)) /\
4633                         (insider_dcprog (SUC foo) ((y,z),r) = x)})
4634                 else
4635                   0) /
4636                (if
4637                   z IN biased_low_states /\ y IN biased_high_states foo
4638                 then
4639                   1 / 32 *
4640                   &
4641                     (CARD
4642                        {r |
4643                         r IN biased_random_states foo /\
4644                         r (STRCAT "coin" (toString 1))}) +
4645                   3 / 32 *
4646                   &
4647                     (CARD
4648                        {r |
4649                         r IN biased_random_states foo /\
4650                         ~r (STRCAT "coin" (toString 1))})
4651                 else
4652                   0))) x
4653        else
4654          0) =
4655        (if
4656          x IN
4657          IMAGE (insider_dcprog (SUC foo))
4658            (biased_high_states foo CROSS biased_low_states CROSS
4659             biased_random_states foo) CROSS
4660          (biased_high_states foo CROSS biased_low_states)
4661        then
4662          (\(x,y,z).
4663             (1 / 32 *
4664                &
4665                  (CARD
4666                     {r |
4667                      r IN biased_random_states foo /\
4668                      r (STRCAT "coin" (toString 1)) /\
4669                      (insider_dcprog (SUC foo) ((y,z),r) = x)}) +
4670                3 / 32 *
4671                &
4672                  (CARD
4673                     {r |
4674                      r IN biased_random_states foo /\
4675                      ~r (STRCAT "coin" (toString 1)) /\
4676                      (insider_dcprog (SUC foo) ((y,z),r) = x)})) *
4677             logr 2
4678               ((1 / 32 *
4679                   &
4680                     (CARD
4681                        {r |
4682                         r IN biased_random_states foo /\
4683                         r (STRCAT "coin" (toString 1)) /\
4684                         (insider_dcprog (SUC foo) ((y,z),r) = x)}) +
4685                   3 / 32 *
4686                   &
4687                     (CARD
4688                        {r |
4689                         r IN biased_random_states foo /\
4690                         ~r (STRCAT "coin" (toString 1)) /\
4691                         (insider_dcprog (SUC foo) ((y,z),r) = x)})) /
4692                (1 / 32 *
4693                   &
4694                     (CARD
4695                        {r |
4696                         r IN biased_random_states foo /\
4697                         r (STRCAT "coin" (toString 1))}) +
4698                   3 / 32 *
4699                   &
4700                     (CARD
4701                        {r |
4702                         r IN biased_random_states foo /\
4703                         ~r (STRCAT "coin" (toString 1))})))) x
4704        else
4705          0)`
4706        by (RW_TAC std_ss [IN_CROSS] >> `x = (FST x, FST(SND x), SND (SND x))` by RW_TAC std_ss [PAIR] >> POP_ORW >> RW_TAC std_ss [])
4707   >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF, IMAGE_FINITE, FINITE_CROSS]
4708   >> Q.UNABBREV_TAC `sumone`
4709   >> RW_TAC real_ss [] >> Q.PAT_X_ASSUM `!f. f IN P` (K ALL_TAC)
4710   >> RW_TAC std_ss [GSYM lg_def]
4711   >> PURE_REWRITE_TAC [ETA_THM]
4712   >> `!(x:bool state) (y:bool state) (z:bool state).
4713        &(CARD
4714                {r |
4715                 r IN biased_random_states foo /\
4716                 r (STRCAT "coin" (toString 1))}) = 2`
4717        by (REPEAT STRIP_TAC
4718            >> `&(CARD
4719                {r |
4720                 r IN biased_random_states foo /\
4721                 r (STRCAT "coin" (toString 1))}) =
4722                REAL_SUM_IMAGE (\r. if r (STRCAT "coin" (toString 1)) then 1 else 0)
4723                (biased_random_states foo)`
4724                by (`FINITE {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}`
4725                    by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4726                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4727                    >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
4728                    >> `FINITE {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4729                    by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4730                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4731                    >>  Q.ABBREV_TAC `s = {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}`
4732                    >> `(biased_random_states foo) =
4733                        {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))} UNION
4734                        {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4735                        by (Q.UNABBREV_TAC `s` >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION] >> DECIDE_TAC)
4736                    >> POP_ORW
4737                    >> Q.UNABBREV_TAC `s`
4738                    >> `DISJOINT {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4739                        {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4740                        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
4741                            >> METIS_TAC [])
4742                    >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
4743                    >> `SIGMA (\r. (if r (STRCAT "coin" (toString 1)) then 1 else 0))
4744                        {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))} = 0`
4745                        by (RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
4746                            >> Suff `!r. (if r IN
4747                                {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))} then
4748                                (if r (STRCAT "coin" (toString 1)) then 1 else 0) else 0) = 0`
4749                            >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
4750                            >> RW_TAC std_ss [GSPECIFICATION])
4751                    >> POP_ORW >> RW_TAC real_ss []
4752                    >> Suff `(\x. (if x IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4753                                then 1 else 0)) =
4754                             (\x. (if x IN {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4755                                then (\r. (if r (STRCAT "coin" (toString 1)) then 1 else 0)) x else 0))`
4756                     >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
4757                         >> METIS_TAC [])
4758                     >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION])
4759            >> POP_ORW
4760            >> Q.UNABBREV_TAC `foo` >> RW_TAC bool_ss [biased_random_states_def]
4761            >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
4762                                                   L_def, SND]
4763                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
4764                                                     THENC (TRY_CONV insider_hl_dup_conv)
4765                                                     THENC (TRY_CONV insider_hl_dup_conv)
4766                                                     THENC (TRY_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
4767                                              THENC EVAL
4768                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
4769                                              THENC EVAL
4770                                              THENC (REPEATC (T_F_UNCHANGED_CONV
4771                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
4772                                                         THENC EVAL)))))))))
4773            >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
4774                         THENC ((ONCE_FIND_CONV ``x DELETE y``
4775                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
4776                                              THENC EVAL
4777                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
4778                                              THENC EVAL
4779                                              THENC (REPEATC (T_F_UNCHANGED_CONV
4780                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
4781                                                         THENC EVAL)))))))
4782                         THENC SIMP_CONV arith_ss []))
4783            >> RW_TAC real_ss [STRCAT_toString_inj])
4784   >> POP_ORW
4785   >> `!(x:bool state) (y:bool state) (z:bool state).
4786        &(CARD
4787                {r |
4788                 r IN biased_random_states foo /\
4789                 ~r (STRCAT "coin" (toString 1))}) = 2`
4790        by (REPEAT STRIP_TAC
4791            >> `&(CARD
4792                {r |
4793                 r IN biased_random_states foo /\
4794                 ~r (STRCAT "coin" (toString 1))}) =
4795                REAL_SUM_IMAGE (\r. if ~r (STRCAT "coin" (toString 1)) then 1 else 0)
4796                (biased_random_states foo)`
4797                by (`FINITE {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}`
4798                    by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4799                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4800                    >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
4801                    >> `FINITE {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4802                    by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
4803                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
4804                    >>  Q.ABBREV_TAC `s = {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4805                    >> `(biased_random_states foo) =
4806                        {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))} UNION
4807                        {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4808                        by (Q.UNABBREV_TAC `s` >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION] >> DECIDE_TAC)
4809                    >> POP_ORW
4810                    >> Q.UNABBREV_TAC `s`
4811                    >> `DISJOINT {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))}
4812                        {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}`
4813                        by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
4814                            >> METIS_TAC [])
4815                    >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
4816                    >> `SIGMA (\r. (if ~r (STRCAT "coin" (toString 1)) then 1 else 0))
4817                        {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))} = 0`
4818                        by (RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
4819                            >> Suff `!r. (if r IN
4820                                {r | r IN biased_random_states foo /\ r (STRCAT "coin" (toString 1))} then
4821                                (if ~r (STRCAT "coin" (toString 1)) then 1 else 0) else 0) = 0`
4822                            >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
4823                            >> RW_TAC std_ss [GSPECIFICATION])
4824                    >> POP_ORW >> RW_TAC real_ss []
4825                    >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
4826                    >> Suff `(\x. (if x IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}
4827                                then 1 else 0)) =
4828                             (\x. (if x IN {r | r IN biased_random_states foo /\ ~r (STRCAT "coin" (toString 1))}
4829                                then (\r. (if ~r (STRCAT "coin" (toString 1)) then 1 else 0)) x else 0))`
4830                     >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
4831                         >> METIS_TAC [])
4832                     >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION])
4833            >> POP_ORW
4834            >> Q.UNABBREV_TAC `foo` >> RW_TAC bool_ss [biased_random_states_def]
4835            >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
4836                                                   L_def, SND]
4837                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
4838                                                     THENC (TRY_CONV insider_hl_dup_conv)
4839                                                     THENC (TRY_CONV insider_hl_dup_conv)
4840                                                     THENC (TRY_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
4841                                              THENC EVAL
4842                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
4843                                              THENC EVAL
4844                                              THENC (REPEATC (T_F_UNCHANGED_CONV
4845                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
4846                                                         THENC EVAL)))))))))
4847            >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
4848                         THENC ((ONCE_FIND_CONV ``x DELETE y``
4849                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
4850                                              THENC EVAL
4851                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
4852                                              THENC EVAL
4853                                              THENC (REPEATC (T_F_UNCHANGED_CONV
4854                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
4855                                                         THENC EVAL)))))))
4856                         THENC SIMP_CONV arith_ss []))
4857            >> RW_TAC real_ss [STRCAT_toString_inj])
4858   >> POP_ORW
4859   >> RW_TAC real_ss []
4860   >> `!(x: bool state) (z:bool state) (a:real) b c1 c2 d. (a * c1  + b * c2) * lg ((a * c1  + b * c2)/d) =
4861                (\c1 c2. (a * c1  + b * c2) * lg ((a * c1  + b * c2)/d)) c1 c2`
4862        by  RW_TAC std_ss []
4863   >> POP_ORW
4864   >> Q.ABBREV_TAC `c1 = (\c1 c2.
4865           (1 / 32 * c1 + 3 / 32 * c2) *
4866           lg ((1 / 32 * c1 + 3 / 32 * c2) / (1 / 2)))`
4867   >> Q.ABBREV_TAC `c2 = (\c1 c2.
4868          (1 / 32 * c1 + 3 / 32 * c2) *
4869          lg ((1 / 32 * c1 + 3 / 32 * c2) / (1 / 4)))`
4870   >> `!(x: bool state) (z:bool state).
4871        &(CARD {(h,r) |
4872                  h IN biased_high_states foo /\
4873                  r IN biased_random_states foo /\
4874                  r (STRCAT "coin" (toString 1)) /\
4875                  (insider_dcprog (SUC foo) ((h,z),r) = x)}) =
4876        REAL_SUM_IMAGE
4877        (\(h,r). if r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)
4878        (biased_high_states foo CROSS biased_random_states foo)`
4879        by (STRIP_TAC >> STRIP_TAC
4880            >> `FINITE {(h,r) |
4881                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4882                        r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}`
4883                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
4884                       >> RW_TAC bool_ss [FINITE_CROSS]
4885                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
4886                       >> POP_ASSUM MP_TAC
4887                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
4888                       >> POP_ORW >> RW_TAC std_ss [])
4889            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
4890            >> `FINITE {(h,r) |
4891                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4892                        ~(r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
4893                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
4894                       >> RW_TAC bool_ss [FINITE_CROSS]
4895                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
4896                       >> POP_ASSUM MP_TAC
4897                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
4898                       >> POP_ORW >> RW_TAC std_ss [])
4899            >> `(biased_high_states foo CROSS biased_random_states foo) =
4900                {(h,r) |
4901                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4902                        r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)} UNION
4903                {(h,r) |
4904                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4905                        ~(r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
4906                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
4907                    >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
4908                    >- (Cases_on `(SND x') (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo)
4909                                        ((FST x',z),SND x') = x)`
4910                        >- (DISJ1_TAC >> Q.EXISTS_TAC `(FST x', SND x')` >> RW_TAC std_ss [PAIR_EQ])
4911                        >> DISJ2_TAC >> Q.EXISTS_TAC `(FST x', SND x')` >> RW_TAC std_ss [PAIR_EQ]
4912                        >> FULL_SIMP_TAC std_ss [DE_MORGAN_THM])
4913                    >> POP_ASSUM MP_TAC
4914                    >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
4915                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ])
4916            >> POP_ORW
4917            >> `DISJOINT {(h,r) |
4918                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4919                        r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}
4920                {(h,r) |
4921                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4922                        ~(r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
4923                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
4924                    >> Cases_on `(SND x') (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo)
4925                                        ((FST x',z),SND x') = x)`
4926                    >- (DISJ2_TAC >> STRIP_TAC
4927                        >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
4928                        >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
4929                    >> DISJ1_TAC >> STRIP_TAC
4930                    >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
4931                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
4932            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
4933            >> `SIGMA (\(h,r). (if r (STRCAT "coin" (toString 1)) /\
4934                        (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0))
4935                  {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
4936                           (~r (STRCAT "coin" (toString 1)) \/
4937                           ~(insider_dcprog (SUC foo) ((h,z),r) = x))} = 0`
4938                by (FULL_SIMP_TAC bool_ss [DE_MORGAN_THM]
4939                    >> RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
4940                    >> Suff `!x'. (if x' IN
4941                        {(h,r) |
4942                         h IN biased_high_states foo /\ r IN biased_random_states foo /\
4943                         (~r (STRCAT "coin" (toString 1)) \/
4944                          ~(insider_dcprog (SUC foo) ((h,z),r) = x))} then
4945                        (\(h,r). (if r (STRCAT "coin" (toString 1)) /\
4946                        (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)) x' else 0) = 0`
4947                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
4948                    >> RW_TAC std_ss [GSPECIFICATION]
4949                    >> POP_ASSUM MP_TAC >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
4950                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4951                    >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
4952                    >> POP_ORW >> RW_TAC std_ss [])
4953            >> POP_ORW >> RW_TAC real_ss []
4954            >> Suff `(\x'. (if x' IN {(h,r) |
4955                h IN biased_high_states foo /\ r IN biased_random_states foo /\
4956                r (STRCAT "coin" (toString 1)) /\
4957                (insider_dcprog (SUC foo) ((h,z),r) = x)}
4958                        then 1 else 0)) =
4959                     (\x'. (if x' IN {(h,r) |
4960                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4961                        r (STRCAT "coin" (toString 1)) /\
4962                        (insider_dcprog (SUC foo) ((h,z),r) = x)}
4963                        then (\(h,r). (if r (STRCAT "coin" (toString 1)) /\
4964                (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)) x' else 0))`
4965             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
4966                 >> METIS_TAC [])
4967             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION]
4968             >> POP_ASSUM MP_TAC >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
4969             >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
4970             >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
4971             >> POP_ORW >> RW_TAC std_ss [])
4972   >> POP_ORW
4973   >> `!(x: bool state) (z:bool state).
4974        &(CARD {(h,r) |
4975                  h IN biased_high_states foo /\
4976                  r IN biased_random_states foo /\
4977                  ~r (STRCAT "coin" (toString 1)) /\
4978                  (insider_dcprog (SUC foo) ((h,z),r) = x)}) =
4979        REAL_SUM_IMAGE
4980        (\(h,r). if ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)
4981        (biased_high_states foo CROSS biased_random_states foo)`
4982        by (STRIP_TAC >> STRIP_TAC
4983            >> `FINITE {(h,r) |
4984                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4985                        ~(~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
4986                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
4987                       >> RW_TAC bool_ss [FINITE_CROSS]
4988                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
4989                       >> POP_ASSUM MP_TAC
4990                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
4991                       >> POP_ORW >> RW_TAC std_ss [])
4992            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
4993            >> `FINITE {(h,r) |
4994                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
4995                        ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}`
4996                   by ((MP_TAC o Q.ISPEC `(biased_high_states foo) CROSS (biased_random_states foo)`) SUBSET_FINITE
4997                       >> RW_TAC bool_ss [FINITE_CROSS]
4998                       >> POP_ASSUM MATCH_MP_TAC >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION, IN_CROSS]
4999                       >> POP_ASSUM MP_TAC
5000                       >> `x'' = (FST x'',SND x'')` by RW_TAC std_ss [PAIR]
5001                       >> POP_ORW >> RW_TAC std_ss [])
5002            >> `(biased_high_states foo CROSS biased_random_states foo) =
5003                {(h,r) |
5004                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
5005                        ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)} UNION
5006                {(h,r) |
5007                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
5008                        ~(~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
5009                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
5010                    >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss [])
5011                    >- (Cases_on `~(SND x') (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((FST x',z),SND x') = x)`
5012                        >- (DISJ1_TAC >> Q.EXISTS_TAC `(FST x', SND x')` >> RW_TAC std_ss [PAIR_EQ])
5013                        >> DISJ2_TAC >> Q.EXISTS_TAC `(FST x', SND x')` >> RW_TAC std_ss [PAIR_EQ]
5014                        >> FULL_SIMP_TAC std_ss [DE_MORGAN_THM])
5015                    >> POP_ASSUM MP_TAC
5016                    >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
5017                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ])
5018            >> POP_ORW
5019            >> `DISJOINT {(h,r) |
5020                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
5021                        ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x)}
5022                {(h,r) |
5023                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
5024                        ~(~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((h,z),r) = x))}`
5025                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION]
5026                    >> Cases_on `~(SND x') (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((FST x',z),SND x') = x)`
5027                    >- (DISJ2_TAC >> STRIP_TAC
5028                        >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
5029                        >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
5030                    >> DISJ1_TAC >> STRIP_TAC
5031                    >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
5032                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ] >> METIS_TAC [PAIR, PAIR_EQ])
5033            >> FULL_SIMP_TAC std_ss [DE_MORGAN_THM]
5034            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
5035            >> `SIGMA (\(h,r). (if ~r (STRCAT "coin" (toString 1)) /\
5036                (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0))
5037                {(h,r) | h IN biased_high_states foo /\ r IN biased_random_states foo /\
5038                (r (STRCAT "coin" (toString 1)) \/ ~(insider_dcprog (SUC foo) ((h,z),r) = x))} = 0`
5039                by (FULL_SIMP_TAC bool_ss [DE_MORGAN_THM]
5040                    >> RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
5041                    >> Suff `!x'. (if x' IN
5042                        {(h,r) |
5043                         h IN biased_high_states foo /\ r IN biased_random_states foo /\
5044                         (r (STRCAT "coin" (toString 1)) \/
5045                         ~(insider_dcprog (SUC foo) ((h,z),r) = x))} then
5046                        (\(h,r). (if ~r (STRCAT "coin" (toString 1)) /\
5047                        (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)) x' else 0) = 0`
5048                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
5049                    >> RW_TAC std_ss [GSPECIFICATION]
5050                    >> POP_ASSUM MP_TAC >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
5051                    >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
5052                    >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
5053                    >> POP_ORW >> RW_TAC std_ss [])
5054            >> POP_ORW >> RW_TAC real_ss []
5055            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
5056            >> Suff `(\x'. (if x' IN {(h,r) |
5057                h IN biased_high_states foo /\ r IN biased_random_states foo /\
5058                ~r (STRCAT "coin" (toString 1)) /\
5059                (insider_dcprog (SUC foo) ((h,z),r) = x)}
5060                        then 1 else 0)) =
5061                     (\x'. (if x' IN {(h,r) |
5062                        h IN biased_high_states foo /\ r IN biased_random_states foo /\
5063                        ~r (STRCAT "coin" (toString 1)) /\
5064                        (insider_dcprog (SUC foo) ((h,z),r) = x)}
5065                        then (\(h,r). (if ~r (STRCAT "coin" (toString 1)) /\
5066                        (insider_dcprog (SUC foo) ((h,z),r) = x) then 1 else 0)) x' else 0))`
5067             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
5068                 >> METIS_TAC [])
5069             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION]
5070             >> POP_ASSUM MP_TAC >> `x'':bool state # bool state = (FST x'', SND x'')` by RW_TAC std_ss [PAIR]
5071             >> POP_ORW >> RW_TAC std_ss [PAIR_EQ]
5072             >> `x':bool state # bool state = (FST x', SND x')` by RW_TAC std_ss [PAIR]
5073             >> POP_ORW >> RW_TAC std_ss [])
5074   >> POP_ORW
5075   >> `!(x: bool state) (y: bool state) (z:bool state).
5076        &(CARD {r |
5077                r IN biased_random_states foo /\
5078                r (STRCAT "coin" (toString 1)) /\
5079                (insider_dcprog (SUC foo) ((y,z),r) = x)}) =
5080        REAL_SUM_IMAGE
5081        (\r. if r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0)
5082        (biased_random_states foo)`
5083        by (STRIP_TAC >> STRIP_TAC >> STRIP_TAC
5084            >> `FINITE {r |
5085                r IN biased_random_states foo /\
5086                r (STRCAT "coin" (toString 1)) /\
5087                (insider_dcprog (SUC foo) ((y,z),r) = x)}`
5088                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
5089                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
5090            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
5091            >> `FINITE {r |
5092                r IN biased_random_states foo /\
5093                ~(r (STRCAT "coin" (toString 1)) /\
5094                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5095                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
5096                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
5097           >> Q.ABBREV_TAC `s = {r | r IN biased_random_states foo /\
5098                                     r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x)}`
5099            >> `(biased_random_states foo) =
5100                {r |
5101                r IN biased_random_states foo /\
5102                r (STRCAT "coin" (toString 1)) /\
5103                (insider_dcprog (SUC foo) ((y,z),r) = x)} UNION
5104                {r |
5105                r IN biased_random_states foo /\
5106                ~(r (STRCAT "coin" (toString 1)) /\
5107                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5108                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
5109                    >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []) >> METIS_TAC [])
5110            >> POP_ORW
5111            >> `DISJOINT {r |
5112                r IN biased_random_states foo /\
5113                r (STRCAT "coin" (toString 1)) /\
5114                (insider_dcprog (SUC foo) ((y,z),r) = x)}
5115                {r |
5116                r IN biased_random_states foo /\
5117                ~(r (STRCAT "coin" (toString 1)) /\
5118                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5119                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> Q.UNABBREV_TAC `s` >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION] >> DECIDE_TAC)
5120            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
5121            >> `SIGMA (\r. (if r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0))
5122                {r | r IN biased_random_states foo /\ (~r (STRCAT "coin" (toString 1)) \/
5123                        ~(insider_dcprog (SUC foo) ((y,z),r) = x))} = 0`
5124                by (FULL_SIMP_TAC bool_ss [DE_MORGAN_THM]
5125                    >> RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
5126                    >> Suff `!r. (if r IN
5127                        {r | r IN biased_random_states foo /\ (~r (STRCAT "coin" (toString 1)) \/
5128                            ~(insider_dcprog (SUC foo) ((y,z),r) = x))} then
5129                        (if r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0) else 0) = 0`
5130                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
5131                    >> RW_TAC std_ss [GSPECIFICATION])
5132            >> POP_ORW >> RW_TAC real_ss []
5133            >> Suff `(\x'. (if x' IN s then 1 else 0)) =
5134                     (\x'. (if x' IN s then (\r. (if r (STRCAT "coin" (toString 1)) /\
5135                                (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0)) x' else 0))`
5136             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
5137                 >> METIS_TAC [])
5138             >> Q.UNABBREV_TAC `s`
5139             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION])
5140   >> POP_ORW
5141   >> `!(x: bool state) (y: bool state) (z:bool state).
5142        &(CARD {r |
5143                r IN biased_random_states foo /\
5144                ~r (STRCAT "coin" (toString 1)) /\
5145                (insider_dcprog (SUC foo) ((y,z),r) = x)}) =
5146        REAL_SUM_IMAGE
5147        (\r. if ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0)
5148        (biased_random_states foo)`
5149        by (STRIP_TAC >> STRIP_TAC >> STRIP_TAC
5150            >> `FINITE {r |
5151                r IN biased_random_states foo /\
5152                ~r (STRCAT "coin" (toString 1)) /\
5153                (insider_dcprog (SUC foo) ((y,z),r) = x)}`
5154                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
5155                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
5156            >> RW_TAC std_ss [GSYM REAL_SUM_IMAGE_EQ_CARD]
5157            >> `FINITE {r |
5158                r IN biased_random_states foo /\
5159                ~(~r (STRCAT "coin" (toString 1)) /\
5160                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5161                   by ((MATCH_MP_TAC o UNDISCH o Q.ISPEC `(biased_random_states foo)`) SUBSET_FINITE
5162                       >> RW_TAC std_ss [SUBSET_DEF, GSPECIFICATION])
5163           >> Q.ABBREV_TAC `s = {r | r IN biased_random_states foo /\
5164                                     ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x)}`
5165            >> `(biased_random_states foo) =
5166                {r |
5167                r IN biased_random_states foo /\
5168                ~r (STRCAT "coin" (toString 1)) /\
5169                (insider_dcprog (SUC foo) ((y,z),r) = x)} UNION
5170                {r |
5171                r IN biased_random_states foo /\
5172                ~(~r (STRCAT "coin" (toString 1)) /\
5173                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5174                by (ONCE_REWRITE_TAC [EXTENSION] >> RW_TAC std_ss [IN_UNION, GSPECIFICATION, IN_CROSS]
5175                    >> (EQ_TAC >> RW_TAC std_ss [] >> RW_TAC std_ss []) >> METIS_TAC [])
5176            >> POP_ORW
5177            >> `DISJOINT {r |
5178                r IN biased_random_states foo /\
5179                ~r (STRCAT "coin" (toString 1)) /\
5180                (insider_dcprog (SUC foo) ((y,z),r) = x)}
5181                {r |
5182                r IN biased_random_states foo /\
5183                ~(~r (STRCAT "coin" (toString 1)) /\
5184                (insider_dcprog (SUC foo) ((y,z),r) = x))}`
5185                by (RW_TAC std_ss [DISJOINT_DEF] >> ONCE_REWRITE_TAC [EXTENSION] >> Q.UNABBREV_TAC `s` >> RW_TAC std_ss [IN_INTER, NOT_IN_EMPTY, GSPECIFICATION] >> DECIDE_TAC)
5186            >> RW_TAC std_ss [REAL_SUM_IMAGE_DISJOINT_UNION] >> POP_ASSUM (K ALL_TAC)
5187            >> `SIGMA (\r. (if ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0))
5188                {r | r IN biased_random_states foo /\ (r (STRCAT "coin" (toString 1)) \/
5189                        ~(insider_dcprog (SUC foo) ((y,z),r) = x))} = 0`
5190                by (FULL_SIMP_TAC bool_ss [DE_MORGAN_THM]
5191                    >> RW_TAC std_ss [Once REAL_SUM_IMAGE_IN_IF]
5192                    >> Suff `!r. (if r IN
5193                        {r | r IN biased_random_states foo /\ (r (STRCAT "coin" (toString 1)) \/
5194                            ~(insider_dcprog (SUC foo) ((y,z),r) = x))} then
5195                        (if ~r (STRCAT "coin" (toString 1)) /\ (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0) else 0) = 0`
5196                    >- RW_TAC std_ss [REAL_SUM_IMAGE_0]
5197                    >> RW_TAC std_ss [GSPECIFICATION])
5198            >> POP_ORW >> RW_TAC real_ss []
5199            >> Suff `(\x'. (if x' IN s then 1 else 0)) =
5200                     (\x'. (if x' IN s then (\r. (if ~r (STRCAT "coin" (toString 1)) /\
5201                                (insider_dcprog (SUC foo) ((y,z),r) = x) then 1 else 0)) x' else 0))`
5202             >- (STRIP_TAC >> POP_ORW >> RW_TAC bool_ss [GSYM REAL_SUM_IMAGE_IN_IF]
5203                 >> METIS_TAC [])
5204             >> Q.UNABBREV_TAC `s`
5205             >> ONCE_REWRITE_TAC [FUN_EQ_THM] >> RW_TAC std_ss [GSPECIFICATION])
5206   >> POP_ORW
5207   >> Q.UNABBREV_TAC `foo`
5208   >> Q.ABBREV_TAC `s1 = (IMAGE (insider_dcprog (SUC (SUC (SUC 0))))
5209         (biased_high_states (SUC (SUC 0)) CROSS biased_low_states CROSS
5210          biased_random_states (SUC (SUC 0))))`
5211   >> Q.ABBREV_TAC `s2 = (s1 CROSS biased_low_states)`
5212   >> Q.ABBREV_TAC `s3 = (s1 CROSS (biased_high_states (SUC (SUC 0)) CROSS biased_low_states))`
5213   >> RW_TAC bool_ss [biased_high_states_def, biased_random_states_def, insider_high_states_set_def,
5214                      insider_dcprog_def, compute_result_alt, insider_set_announcements_alt, XOR_announces_def,
5215                      H_def, L_def, PAIR, ETA_THM, xor_def, STRCAT_toString_inj]
5216   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
5217                                                   L_def, SND]
5218                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
5219                                                     THENC (TRY_CONV insider_hl_dup_conv)
5220                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
5221   >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
5222                         THENC ((ONCE_FIND_CONV ``x DELETE y``
5223                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
5224                                              THENC EVAL
5225                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
5226                                              THENC EVAL
5227                                              THENC (REPEATC (T_F_UNCHANGED_CONV
5228                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
5229                                                         THENC EVAL)))))))
5230                         THENC SIMP_CONV arith_ss []))
5231   >> RW_TAC real_ss [STRCAT_toString_inj]
5232   >> Q.UNABBREV_TAC `s3` >> Q.UNABBREV_TAC `s2` >> Q.UNABBREV_TAC `s1`
5233   >> RW_TAC bool_ss [biased_high_states_def, biased_low_states_def, biased_random_states_def, insider_high_states_set_def,
5234                      insider_dcprog_def, compute_result_alt, insider_set_announcements_alt, XOR_announces_def,
5235                      H_def, L_def, PAIR, ETA_THM, xor_def, STRCAT_toString_inj]
5236   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
5237                                                   L_def, SND]
5238                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
5239                                                     THENC (TRY_CONV insider_hl_dup_conv)
5240                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
5241   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
5242                                                   L_def, SND]
5243                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
5244                                                     THENC (TRY_CONV insider_hl_dup_conv)
5245                                                     THENC (TRY_CONV insider_hl_dup_conv)))))
5246   >> RW_TAC std_ss []
5247   >> CONV_TAC (SIMP_CONV bool_ss [CROSS_EQNS, PAIR_EQ, IMAGE_UNION, IMAGE_INSERT, IMAGE_EMPTY, UNION_EMPTY, FST,
5248                                                   L_def, SND]
5249                                THENC (FIND_CONV ``x UNION y`` (UNION_CONV (SIMP_CONV bool_ss [PAIR_EQ]
5250                                                     THENC (TRY_CONV insider_hl_dup_conv)
5251                                                     THENC (TRY_CONV insider_hl_dup_conv)
5252                                                     THENC (TRY_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
5253                                              THENC EVAL
5254                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
5255                                              THENC EVAL
5256                                              THENC (REPEATC (T_F_UNCHANGED_CONV
5257                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
5258                                                         THENC EVAL)))))))))
5259   >> RW_TAC real_ss [STRCAT_toString_inj]
5260  >> CONV_TAC (REPEATC (SIMP_CONV set_ss [REAL_SUM_IMAGE_THM]
5261                         THENC ((ONCE_FIND_CONV ``x DELETE y``
5262                                (DELETE_CONV (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
5263                                              THENC EVAL
5264                                              THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
5265                                              THENC EVAL
5266                                              THENC (REPEATC (T_F_UNCHANGED_CONV
5267                                                        (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
5268                                                         THENC EVAL)))))))
5269                         THENC SIMP_CONV arith_ss []))
5270   >> SIMP_TAC real_ss [STRCAT_toString_inj]
5271   >> CONV_TAC (ONCE_FIND_CONV ``if (x=y) then (1:real) else 0`` (RATOR_CONV (RATOR_CONV (RAND_CONV
5272        (SIMP_CONV arith_ss [FUN_EQ_THM, PAIR_EQ, COND_EXPAND, EQ_IMP_THM, FORALL_AND_THM, DISJ_IMP_THM]
5273         THENC EVAL
5274         THENC SIMP_CONV bool_ss [COND_EXPAND, LEFT_AND_OVER_OR, FORALL_AND_THM, DISJ_IMP_THM]
5275         THENC EVAL
5276         THENC (REPEATC (T_F_UNCHANGED_CONV
5277         (SIMP_CONV bool_ss [fun_eq_lem6, GSYM LEFT_AND_OVER_OR, FORALL_AND_THM]
5278         THENC EVAL))))))
5279                                                                  THENC SIMP_CONV bool_ss []))
5280   >> RW_TAC real_ss []
5281   >> Q.UNABBREV_TAC `c2` >> Q.UNABBREV_TAC `c1`
5282   >> RW_TAC real_ss []
5283   >> `lg(1/4) = ~2` by (RW_TAC real_ss [GSYM REAL_INV_1OVER, lg_inv] >> `4 = 2 pow 2` by RW_TAC real_ss [pow] >> POP_ORW >> RW_TAC std_ss [lg_pow])
5284   >> POP_ORW
5285   >> RW_TAC real_ss []
5286   >> RW_TAC std_ss [real_div]
5287   >> RW_TAC real_ss [lg_mul, REAL_INV_POS, lg_inv]
5288   >> `lg(8) = 3` by (`8 = 2 pow 3` by RW_TAC real_ss [pow] >> POP_ORW >> RW_TAC std_ss [lg_pow])
5289   >> POP_ORW >> RW_TAC real_ss [GSYM real_sub, REAL_SUB_LDISTRIB, lg_1]
5290   >> RW_TAC real_ss [REAL_ADD_ASSOC, GSYM real_sub, REAL_ARITH ``!x (y:real). x - y - y = x - 2 * y``,
5291                      REAL_ARITH ``!x (y:real) c. x - c * y - y = x - (c + 1) * y``,
5292                      REAL_ARITH ``!x (y:real) z. x - y + z = x + z - y``]
5293   >> RW_TAC std_ss [GSYM REAL_ADD_ASSOC, REAL_ARITH ``!(y:real). y + y = 2 * y``,
5294                     REAL_ARITH ``!(y:real) c. y + c * y = (c + 1) * y``]
5295   >> RW_TAC real_ss [REAL_SUB_LDISTRIB, REAL_MUL_ASSOC, REAL_INV_1OVER,
5296                      REAL_ARITH ``!(x:real) y z a. x + (y - z) - a = y + (x - (z + a))``, GSYM real_sub]
5297   >> Q.UNABBREV_TAC `v` >> RW_TAC std_ss []);
5298val _ = export_theory ();
5299