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