| author | wenzelm | 
| Fri, 27 Jan 2006 19:03:19 +0100 | |
| changeset 18812 | a4554848b59e | 
| parent 16417 | 9bc16273c2d4 | 
| child 20768 | 1d478c2d621f | 
| permissions | -rw-r--r-- | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 1 | (* Title: HOL/Auth/ZhouGollmann | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 2 | ID: $Id$ | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 3 | Author: Giampaolo Bella and L C Paulson, Cambridge Univ Computer Lab | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 4 | Copyright 2003 University of Cambridge | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 5 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 6 | The protocol of | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 7 | Jianying Zhou and Dieter Gollmann, | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 8 | A Fair Non-Repudiation Protocol, | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 9 | Security and Privacy 1996 (Oakland) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 10 | 55-61 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 11 | *) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 12 | |
| 16417 | 13 | theory ZhouGollmann imports Public begin | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 14 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 15 | syntax | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 16 | TTP :: agent | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 17 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 18 | translations | 
| 14741 | 19 | "TTP" == " Server " | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 20 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 21 | syntax | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 22 | f_sub :: nat | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 23 | f_nro :: nat | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 24 | f_nrr :: nat | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 25 | f_con :: nat | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 26 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 27 | translations | 
| 14741 | 28 | "f_sub" == " 5 " | 
| 29 | "f_nro" == " 2 " | |
| 30 | "f_nrr" == " 3 " | |
| 31 | "f_con" == " 4 " | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 32 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 33 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 34 | constdefs | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 35 | broken :: "agent set" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 36 |     --{*the compromised honest agents; TTP is included as it's not allowed to
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 37 | use the protocol*} | 
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 38 |    "broken == bad - {Spy}"
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 39 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 40 | declare broken_def [simp] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 41 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 42 | consts zg :: "event list set" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 43 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 44 | inductive zg | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 45 | intros | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 46 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 47 | Nil: "[] \<in> zg" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 48 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 49 | Fake: "[| evsf \<in> zg; X \<in> synth (analz (spies evsf)) |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 50 | ==> Says Spy B X # evsf \<in> zg" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 51 | |
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 52 | Reception: "[| evsr \<in> zg; Says A B X \<in> set evsr |] ==> Gets B X # evsr \<in> zg" | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 53 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 54 | (*L is fresh for honest agents. | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 55 | We don't require K to be fresh because we don't bother to prove secrecy! | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 56 | We just assume that the protocol's objective is to deliver K fairly, | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 57 | rather than to keep M secret.*) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 58 | ZG1: "[| evs1 \<in> zg; Nonce L \<notin> used evs1; C = Crypt K (Number m); | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 59 | K \<in> symKeys; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 60 | 	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|}|]
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 61 |        ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} # evs1 \<in> zg"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 62 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 63 | (*B must check that NRO is A's signature to learn the sender's name*) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 64 | ZG2: "[| evs2 \<in> zg; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 65 | 	   Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs2;
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 66 | 	   NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 67 | 	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|}|]
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 68 |        ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} # evs2  \<in>  zg"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 69 | |
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 70 | (*A must check that NRR is B's signature to learn the sender's name; | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 71 | without spy, the matching label would be enough*) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 72 | ZG3: "[| evs3 \<in> zg; C = Crypt K M; K \<in> symKeys; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 73 | 	   Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs3;
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 74 | 	   Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs3;
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 75 | 	   NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 76 | 	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|}|]
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 77 |        ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 78 | # evs3 \<in> zg" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 79 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 80 | (*TTP checks that sub_K is A's signature to learn who issued K, then | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 81 | gives credentials to A and B. The Notes event models the availability of | 
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 82 | the credentials, but the act of fetching them is not modelled. We also | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 83 | give con_K to the Spy. This makes the threat model more dangerous, while | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 84 |    also allowing lemma @{text Crypt_used_imp_spies} to omit the condition
 | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 85 |    @{term "K \<noteq> priK TTP"}. *)
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 86 | ZG4: "[| evs4 \<in> zg; K \<in> symKeys; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 87 | 	   Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 88 | \<in> set evs4; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 89 | 	   sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 90 | 	   con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 91 | Nonce L, Key K|}|] | 
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 92 | ==> Says TTP Spy con_K | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 93 | # | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 94 | 	   Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
 | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 95 | # evs4 \<in> zg" | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 96 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 97 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 98 | declare Says_imp_knows_Spy [THEN analz.Inj, dest] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 99 | declare Fake_parts_insert_in_Un [dest] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 100 | declare analz_into_parts [dest] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 101 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 102 | declare symKey_neq_priEK [simp] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 103 | declare symKey_neq_priEK [THEN not_sym, simp] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 104 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 105 | |
| 14146 | 106 | text{*A "possibility property": there are traces that reach the end*}
 | 
| 107 | lemma "[|A \<noteq> B; TTP \<noteq> A; TTP \<noteq> B; K \<in> symKeys|] ==> | |
| 108 | \<exists>L. \<exists>evs \<in> zg. | |
| 109 |            Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K,
 | |
| 110 |                Crypt (priK TTP) {|Number f_con, Agent A, Agent B, Nonce L, Key K|} |}
 | |
| 111 | \<in> set evs" | |
| 112 | apply (intro exI bexI) | |
| 113 | apply (rule_tac [2] zg.Nil | |
| 114 | [THEN zg.ZG1, THEN zg.Reception [of _ A B], | |
| 115 | THEN zg.ZG2, THEN zg.Reception [of _ B A], | |
| 116 | THEN zg.ZG3, THEN zg.Reception [of _ A TTP], | |
| 117 | THEN zg.ZG4]) | |
| 118 | apply (possibility, auto) | |
| 119 | done | |
| 120 | ||
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 121 | subsection {*Basic Lemmas*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 122 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 123 | lemma Gets_imp_Says: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 124 | "[| Gets B X \<in> set evs; evs \<in> zg |] ==> \<exists>A. Says A B X \<in> set evs" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 125 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 126 | apply (erule zg.induct, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 127 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 128 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 129 | lemma Gets_imp_knows_Spy: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 130 | "[| Gets B X \<in> set evs; evs \<in> zg |] ==> X \<in> spies evs" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 131 | by (blast dest!: Gets_imp_Says Says_imp_knows_Spy) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 132 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 133 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 134 | text{*Lets us replace proofs about @{term "used evs"} by simpler proofs 
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 135 | about @{term "parts (spies evs)"}.*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 136 | lemma Crypt_used_imp_spies: | 
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 137 | "[| Crypt K X \<in> used evs; evs \<in> zg |] | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 138 | ==> Crypt K X \<in> parts (spies evs)" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 139 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 140 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 141 | apply (simp_all add: parts_insert_knows_A) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 142 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 143 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 144 | lemma Notes_TTP_imp_Gets: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 145 |      "[|Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K |}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 146 | \<in> set evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 147 |         sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 148 | evs \<in> zg|] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 149 |     ==> Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 150 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 151 | apply (erule zg.induct, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 152 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 153 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 154 | text{*For reasoning about C, which is encrypted in message ZG2*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 155 | lemma ZG2_msg_in_parts_spies: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 156 |      "[|Gets B {|F, B', L, C, X|} \<in> set evs; evs \<in> zg|]
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 157 | ==> C \<in> parts (spies evs)" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 158 | by (blast dest: Gets_imp_Says) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 159 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 160 | (*classical regularity lemma on priK*) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 161 | lemma Spy_see_priK [simp]: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 162 | "evs \<in> zg ==> (Key (priK A) \<in> parts (spies evs)) = (A \<in> bad)" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 163 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 164 | apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 165 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 166 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 167 | text{*So that blast can use it too*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 168 | declare Spy_see_priK [THEN [2] rev_iffD1, dest!] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 169 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 170 | lemma Spy_analz_priK [simp]: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 171 | "evs \<in> zg ==> (Key (priK A) \<in> analz (spies evs)) = (A \<in> bad)" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 172 | by auto | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 173 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 174 | |
| 14741 | 175 | subsection{*About NRO: Validity for @{term B}*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 176 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 177 | text{*Below we prove that if @{term NRO} exists then @{term A} definitely
 | 
| 14741 | 178 | sent it, provided @{term A} is not broken.*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 179 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 180 | text{*Strong conclusion for a good agent*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 181 | lemma NRO_validity_good: | 
| 14741 | 182 |      "[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
 | 
| 183 | NRO \<in> parts (spies evs); | |
| 184 | A \<notin> bad; evs \<in> zg |] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 185 |      ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 186 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 187 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 188 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 189 | apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 190 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 191 | |
| 14741 | 192 | lemma NRO_sender: | 
| 193 |      "[|Says A' B {|n, b, l, C, Crypt (priK A) X|} \<in> set evs; evs \<in> zg|]
 | |
| 194 |     ==> A' \<in> {A,Spy}"
 | |
| 195 | apply (erule rev_mp) | |
| 196 | apply (erule zg.induct, simp_all) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 197 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 198 | |
| 14741 | 199 | text{*Holds also for @{term "A = Spy"}!*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 200 | theorem NRO_validity: | 
| 15047 | 201 |      "[|Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs;
 | 
| 14741 | 202 |         NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
 | 
| 203 | A \<notin> broken; evs \<in> zg |] | |
| 204 |      ==> Says A B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs"
 | |
| 15047 | 205 | apply (drule Gets_imp_Says, assumption) | 
| 14741 | 206 | apply clarify | 
| 207 | apply (frule NRO_sender, auto) | |
| 15047 | 208 | txt{*We are left with the case where the sender is @{term Spy} and not
 | 
| 209 |   equal to @{term A}, because @{term "A \<notin> bad"}. 
 | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 210 |   Thus theorem @{text NRO_validity_good} applies.*}
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 211 | apply (blast dest: NRO_validity_good [OF refl]) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 212 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 213 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 214 | |
| 14741 | 215 | subsection{*About NRR: Validity for @{term A}*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 216 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 217 | text{*Below we prove that if @{term NRR} exists then @{term B} definitely
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 218 | sent it, provided @{term B} is not broken.*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 219 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 220 | text{*Strong conclusion for a good agent*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 221 | lemma NRR_validity_good: | 
| 14741 | 222 |      "[|NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
 | 
| 223 | NRR \<in> parts (spies evs); | |
| 224 | B \<notin> bad; evs \<in> zg |] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 225 |      ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 226 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 227 | apply (erule rev_mp) | 
| 14741 | 228 | apply (erule zg.induct) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 229 | apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 230 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 231 | |
| 14741 | 232 | lemma NRR_sender: | 
| 233 |      "[|Says B' A {|n, a, l, Crypt (priK B) X|} \<in> set evs; evs \<in> zg|]
 | |
| 234 |     ==> B' \<in> {B,Spy}"
 | |
| 235 | apply (erule rev_mp) | |
| 236 | apply (erule zg.induct, simp_all) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 237 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 238 | |
| 14741 | 239 | text{*Holds also for @{term "B = Spy"}!*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 240 | theorem NRR_validity: | 
| 14741 | 241 |      "[|Says B' A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs;
 | 
| 242 |         NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
 | |
| 243 | B \<notin> broken; evs \<in> zg|] | |
| 244 |     ==> Says B A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs"
 | |
| 245 | apply clarify | |
| 246 | apply (frule NRR_sender, auto) | |
| 247 | txt{*We are left with the case where @{term "B' = Spy"} and  @{term "B' \<noteq> B"},
 | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 248 |   i.e. @{term "B \<notin> bad"}, when we can apply @{text NRR_validity_good}.*}
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 249 | apply (blast dest: NRR_validity_good [OF refl]) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 250 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 251 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 252 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 253 | subsection{*Proofs About @{term sub_K}*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 254 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 255 | text{*Below we prove that if @{term sub_K} exists then @{term A} definitely
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 256 | sent it, provided @{term A} is not broken.  *}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 257 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 258 | text{*Strong conclusion for a good agent*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 259 | lemma sub_K_validity_good: | 
| 14741 | 260 |      "[|sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 261 | sub_K \<in> parts (spies evs); | |
| 262 | A \<notin> bad; evs \<in> zg |] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 263 |      ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs"
 | 
| 14741 | 264 | apply clarify | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 265 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 266 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 267 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 268 | txt{*Fake*} 
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 269 | apply (blast dest!: Fake_parts_sing_imp_Un) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 270 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 271 | |
| 14741 | 272 | lemma sub_K_sender: | 
| 273 |      "[|Says A' TTP {|n, b, l, k, Crypt (priK A) X|} \<in> set evs;  evs \<in> zg|]
 | |
| 274 |     ==> A' \<in> {A,Spy}"
 | |
| 275 | apply (erule rev_mp) | |
| 276 | apply (erule zg.induct, simp_all) | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 277 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 278 | |
| 14741 | 279 | text{*Holds also for @{term "A = Spy"}!*}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 280 | theorem sub_K_validity: | 
| 15047 | 281 |      "[|Gets TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs;
 | 
| 14741 | 282 |         sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 283 | A \<notin> broken; evs \<in> zg |] | |
| 284 |      ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs"
 | |
| 15047 | 285 | apply (drule Gets_imp_Says, assumption) | 
| 14741 | 286 | apply clarify | 
| 287 | apply (frule sub_K_sender, auto) | |
| 15047 | 288 | txt{*We are left with the case where the sender is @{term Spy} and not
 | 
| 289 |   equal to @{term A}, because @{term "A \<notin> bad"}. 
 | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 290 |   Thus theorem @{text sub_K_validity_good} applies.*}
 | 
| 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 291 | apply (blast dest: sub_K_validity_good [OF refl]) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 292 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 293 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 294 | |
| 14741 | 295 | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 296 | subsection{*Proofs About @{term con_K}*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 297 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 298 | text{*Below we prove that if @{term con_K} exists, then @{term TTP} has it,
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 299 | and therefore @{term A} and @{term B}) can get it too.  Moreover, we know
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 300 | that @{term A} sent @{term sub_K}*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 301 | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 302 | lemma con_K_validity: | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 303 | "[|con_K \<in> used evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 304 | con_K = Crypt (priK TTP) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 305 |                   {|Number f_con, Agent A, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 306 | evs \<in> zg |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 307 |     ==> Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 308 | \<in> set evs" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 309 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 310 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 311 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 312 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 313 | txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 314 | apply (blast dest!: Fake_parts_sing_imp_Un) | 
| 14741 | 315 | txt{*ZG2*} 
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 316 | apply (blast dest: parts_cut) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 317 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 318 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 319 | text{*If @{term TTP} holds @{term con_K} then @{term A} sent
 | 
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 320 |  @{term sub_K}.  We assume that @{term A} is not broken.  Importantly, nothing
 | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 321 |   needs to be assumed about the form of @{term con_K}!*}
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 322 | lemma Notes_TTP_imp_Says_A: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 323 |      "[|Notes TTP {|Number f_con, Agent A, Agent B, Nonce L, Key K, con_K|}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 324 | \<in> set evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 325 |         sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 326 | A \<notin> broken; evs \<in> zg|] | 
| 14741 | 327 |      ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs"
 | 
| 328 | apply clarify | |
| 329 | apply (erule rev_mp) | |
| 330 | apply (erule zg.induct) | |
| 331 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | |
| 332 | txt{*ZG4*}
 | |
| 15047 | 333 | apply clarify | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 334 | apply (rule sub_K_validity, auto) | 
| 14741 | 335 | done | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 336 | |
| 14736 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 337 | text{*If @{term con_K} exists, then @{term A} sent @{term sub_K}.  We again
 | 
| 
7104394df99a
broken no longer includes TTP, and other minor changes
 paulson parents: 
14207diff
changeset | 338 |    assume that @{term A} is not broken. *}
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 339 | theorem B_sub_K_validity: | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 340 | "[|con_K \<in> used evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 341 |         con_K = Crypt (priK TTP) {|Number f_con, Agent A, Agent B,
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 342 | Nonce L, Key K|}; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 343 |         sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 14741 | 344 | A \<notin> broken; evs \<in> zg|] | 
| 345 |      ==> Says A TTP {|Number f_sub, Agent B, Nonce L, Key K, sub_K|} \<in> set evs"
 | |
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 346 | by (blast dest: con_K_validity Notes_TTP_imp_Says_A) | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 347 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 348 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 349 | subsection{*Proving fairness*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 350 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 351 | text{*Cannot prove that, if @{term B} has NRO, then  @{term A} has her NRR.
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 352 | It would appear that @{term B} has a small advantage, though it is
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 353 | useless to win disputes: @{term B} needs to present @{term con_K} as well.*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 354 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 355 | text{*Strange: unicity of the label protects @{term A}?*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 356 | lemma A_unicity: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 357 |      "[|NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 358 | NRO \<in> parts (spies evs); | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 359 |         Says A B {|Number f_nro, Agent B, Nonce L, Crypt K M', NRO'|}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 360 | \<in> set evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 361 | A \<notin> bad; evs \<in> zg |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 362 | ==> M'=M" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 363 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 364 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 365 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 366 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 367 | apply (frule_tac [5] ZG2_msg_in_parts_spies, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 368 | txt{*ZG1: freshness*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 369 | apply (blast dest: parts.Body) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 370 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 371 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 372 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 373 | text{*Fairness lemma: if @{term sub_K} exists, then @{term A} holds 
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 374 | NRR. Relies on unicity of labels.*} | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 375 | lemma sub_K_implies_NRR: | 
| 14741 | 376 |      "[| NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|};
 | 
| 377 |          NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, Crypt K M|};
 | |
| 378 | sub_K \<in> parts (spies evs); | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 379 | NRO \<in> parts (spies evs); | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 380 |          sub_K = Crypt (priK A) {|Number f_sub, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 381 | A \<notin> bad; evs \<in> zg |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 382 |      ==> Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 383 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 384 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 385 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 386 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 387 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 388 | txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 389 | apply blast | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 390 | txt{*ZG1: freshness*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 391 | apply (blast dest: parts.Body) | 
| 14741 | 392 | txt{*ZG3*} 
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 393 | apply (blast dest: A_unicity [OF refl]) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 394 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 395 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 396 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 397 | lemma Crypt_used_imp_L_used: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 398 |      "[| Crypt (priK TTP) {|F, A, B, L, K|} \<in> used evs; evs \<in> zg |]
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 399 | ==> L \<in> used evs" | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 400 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 401 | apply (erule zg.induct, auto) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 402 | txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 403 | apply (blast dest!: Fake_parts_sing_imp_Un) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 404 | txt{*ZG2: freshness*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 405 | apply (blast dest: parts.Body) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 406 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 407 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 408 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 409 | text{*Fairness for @{term A}: if @{term con_K} and @{term NRO} exist, 
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 410 | then @{term A} holds NRR.  @{term A} must be uncompromised, but there is no
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 411 | assumption about @{term B}.*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 412 | theorem A_fairness_NRO: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 413 | "[|con_K \<in> used evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 414 | NRO \<in> parts (spies evs); | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 415 | con_K = Crypt (priK TTP) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 416 |                       {|Number f_con, Agent A, Agent B, Nonce L, Key K|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 417 |         NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, Crypt K M|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 418 |         NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, Crypt K M|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 419 | A \<notin> bad; evs \<in> zg |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 420 |     ==> Gets A {|Number f_nrr, Agent A, Nonce L, NRR|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 421 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 422 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 423 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 424 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 425 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 426 |    txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 427 | apply (simp add: parts_insert_knows_A) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 428 | apply (blast dest: Fake_parts_sing_imp_Un) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 429 |   txt{*ZG1*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 430 | apply (blast dest: Crypt_used_imp_L_used) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 431 |  txt{*ZG2*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 432 | apply (blast dest: parts_cut) | 
| 14741 | 433 | txt{*ZG4*} 
 | 
| 434 | apply (blast intro: sub_K_implies_NRR [OF refl] | |
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 435 | dest: Gets_imp_knows_Spy [THEN parts.Inj]) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 436 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 437 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 438 | text{*Fairness for @{term B}: NRR exists at all, then @{term B} holds NRO.
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 439 | @{term B} must be uncompromised, but there is no assumption about @{term
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 440 | A}. *} | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 441 | theorem B_fairness_NRR: | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 442 | "[|NRR \<in> used evs; | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 443 |         NRR = Crypt (priK B) {|Number f_nrr, Agent A, Nonce L, C|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 444 |         NRO = Crypt (priK A) {|Number f_nro, Agent B, Nonce L, C|};
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 445 | B \<notin> bad; evs \<in> zg |] | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 446 |     ==> Gets B {|Number f_nro, Agent B, Nonce L, C, NRO|} \<in> set evs"
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 447 | apply clarify | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 448 | apply (erule rev_mp) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 449 | apply (erule zg.induct) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 450 | apply (frule_tac [5] ZG2_msg_in_parts_spies, simp_all) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 451 | txt{*Fake*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 452 | apply (blast dest!: Fake_parts_sing_imp_Un) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 453 | txt{*ZG2*}
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 454 | apply (blast dest: parts_cut) | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 455 | done | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 456 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 457 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 458 | text{*If @{term con_K} exists at all, then @{term B} can get it, by @{text
 | 
| 15068 
58d216b32199
minor tweaks to go with the new version of the Accountability paper
 paulson parents: 
15047diff
changeset | 459 | con_K_validity}.  Cannot conclude that also NRO is available to @{term B},
 | 
| 14145 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 460 | because if @{term A} were unfair, @{term A} could build message 3 without
 | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 461 | building message 1, which contains NRO. *} | 
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 462 | |
| 
2e31b8cc8788
ZhouGollmann: new example (fair non-repudiation protocol)
 paulson parents: diff
changeset | 463 | end |