src/HOL/SET_Protocol/Public_SET.thy
author paulson <lp15@cam.ac.uk>
Thu, 04 Apr 2019 14:19:33 +0100
changeset 70044 da5857dbcbb9
parent 69597 ff784d5a5bfb
child 82695 d93ead9ac6df
permissions -rw-r--r--
More group theory. Sum and product indexed by the non-neutral part of a set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     1
(*  Title:      HOL/SET_Protocol/Public_SET.thy
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
     2
    Author:     Giampaolo Bella
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
     3
    Author:     Fabio Massacci
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
     4
    Author:     Lawrence C Paulson
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     5
*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     6
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
     7
section\<open>The Public-Key Theory, Modified for SET\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
     8
33028
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
     9
theory Public_SET
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    10
imports Event_SET
9aa8bfb1649d modernized session SET_Protocol;
wenzelm
parents: 32960
diff changeset
    11
begin
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    12
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    13
subsection\<open>Symmetric and Asymmetric Keys\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    14
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    15
text\<open>definitions influenced by the wish to assign asymmetric keys 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    16
  - since the beginning - only to RCA and CAs, namely we need a partial 
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    17
  function on type Agent\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    18
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    19
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    20
text\<open>The SET specs mention two signature keys for CAs - we only have one\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    21
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    22
consts
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    23
  publicKey :: "[bool, agent] \<Rightarrow> key"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
    24
    \<comment> \<open>the boolean is TRUE if a signing key\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    25
35068
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
    26
abbreviation "pubEK == publicKey False"
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
    27
abbreviation "pubSK == publicKey True"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    28
35068
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
    29
(*BEWARE!! priEK, priSK DON'T WORK with inj, range, image, etc.*)
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
    30
abbreviation "priEK A == invKey (pubEK A)"
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
    31
abbreviation "priSK A == invKey (pubSK A)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    32
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    33
text\<open>By freeness of agents, no two agents have the same key. Since
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    34
 \<^term>\<open>True\<noteq>False\<close>, no agent has the same signing and encryption keys.\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    35
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    36
specification (publicKey)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    37
  injective_publicKey:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    38
    "publicKey b A = publicKey c A' \<Longrightarrow> b=c \<and> A=A'"
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
    39
(*<*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    40
   apply (rule exI [of _ "%b A. 2 * nat_of_agent A + (if b then 1 else 0)"]) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    41
   apply (auto simp add: inj_on_def inj_nat_of_agent [THEN inj_eq] split: agent.split) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    42
   apply (drule_tac f="%x. x mod 2" in arg_cong, simp add: mod_Suc)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    43
(*or this, but presburger won't abstract out the function applications
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    44
   apply presburger+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    45
*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    46
   done                       
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
    47
(*>*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    48
45827
66c68453455c modernized specifications;
wenzelm
parents: 39758
diff changeset
    49
axiomatization where
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    50
  (*No private key equals any public key (essential to ensure that private
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    51
    keys are private!) *)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    52
  privateKey_neq_publicKey [iff]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    53
      "invKey (publicKey b A) \<noteq> publicKey b' A'"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    54
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    55
declare privateKey_neq_publicKey [THEN not_sym, iff]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    56
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    57
  
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    58
subsection\<open>Initial Knowledge\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    59
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    60
text\<open>This information is not necessary.  Each protocol distributes any needed
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    61
certificates, and anyway our proofs require a formalization of the Spy's 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    62
knowledge only.  However, the initial knowledge is as follows:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    63
   All agents know RCA's public keys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    64
   RCA and CAs know their own respective keys;
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    65
   RCA (has already certified and therefore) knows all CAs public keys; 
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
    66
   Spy knows all keys of all bad agents.\<close>
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    67
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    68
overloading initState \<equiv> "initState"
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    69
begin
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    70
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    71
primrec initState where
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
    72
(*<*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    73
  initState_CA:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    74
    "initState (CA i)  =
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    75
       (if i=0 then Key ` ({priEK RCA, priSK RCA} \<union>
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
    76
                            pubEK ` (range CA) \<union> pubSK ` (range CA))
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    77
        else {Key (priEK (CA i)), Key (priSK (CA i)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    78
              Key (pubEK (CA i)), Key (pubSK (CA i)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    79
              Key (pubEK RCA), Key (pubSK RCA)})"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    80
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    81
| initState_Cardholder:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    82
    "initState (Cardholder i)  =    
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    83
       {Key (priEK (Cardholder i)), Key (priSK (Cardholder i)),
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    84
        Key (pubEK (Cardholder i)), Key (pubSK (Cardholder i)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    85
        Key (pubEK RCA), Key (pubSK RCA)}"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    86
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    87
| initState_Merchant:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    88
    "initState (Merchant i)  =    
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    89
       {Key (priEK (Merchant i)), Key (priSK (Merchant i)),
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    90
        Key (pubEK (Merchant i)), Key (pubSK (Merchant i)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    91
        Key (pubEK RCA), Key (pubSK RCA)}"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    92
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    93
| initState_PG:
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    94
    "initState (PG i)  = 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
    95
       {Key (priEK (PG i)), Key (priSK (PG i)),
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    96
        Key (pubEK (PG i)), Key (pubSK (PG i)),
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
    97
        Key (pubEK RCA), Key (pubSK RCA)}"
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
    98
(*>*)
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
    99
| initState_Spy:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   100
    "initState Spy = Key ` (invKey ` pubEK ` bad \<union>
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   101
                            invKey ` pubSK ` bad \<union>
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   102
                            range pubEK \<union> range pubSK)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   103
39758
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
   104
end
b8a53e3a0ee2 modernized primrecs
haftmann
parents: 36866
diff changeset
   105
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   106
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   107
text\<open>Injective mapping from agents to PANs: an agent can have only one card\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   108
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   109
consts  pan :: "agent \<Rightarrow> nat"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   110
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   111
specification (pan)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   112
  inj_pan: "inj pan"
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   113
  \<comment> \<open>No two agents have the same PAN\<close>
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
   114
(*<*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   115
   apply (rule exI [of _ "nat_of_agent"]) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   116
   apply (simp add: inj_on_def inj_nat_of_agent [THEN inj_eq]) 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   117
   done
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
   118
(*>*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   119
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   120
declare inj_pan [THEN inj_eq, iff]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   121
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   122
consts
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   123
  XOR :: "nat*nat \<Rightarrow> nat"  \<comment> \<open>no properties are assumed of exclusive-or\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   124
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   125
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   126
subsection\<open>Signature Primitives\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   127
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   128
definition
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   129
 (* Signature = Message + signed Digest *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   130
  sign :: "[key, msg]\<Rightarrow>msg"
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   131
  where "sign K X = \<lbrace>X, Crypt K (Hash X) \<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   132
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   133
definition
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   134
 (* Signature Only = signed Digest Only *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   135
  signOnly :: "[key, msg]\<Rightarrow>msg"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   136
  where "signOnly K X = Crypt K (Hash X)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   137
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   138
definition
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   139
 (* Signature for Certificates = Message + signed Message *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   140
  signCert :: "[key, msg]\<Rightarrow>msg"
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   141
  where "signCert K X = \<lbrace>X, Crypt K X \<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   142
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   143
definition
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   144
 (* Certification Authority's Certificate.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   145
    Contains agent name, a key, a number specifying the key's target use,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   146
              a key to sign the entire certificate.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   147
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   148
    Should prove if signK=priSK RCA and C=CA i,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   149
                  then Ka=pubEK i or pubSK i depending on T  ??
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   150
 *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   151
  cert :: "[agent, key, msg, key] \<Rightarrow> msg"
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   152
  where "cert A Ka T signK = signCert signK \<lbrace>Agent A, Key Ka, T\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   153
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   154
definition
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   155
 (* Cardholder's Certificate.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   156
    Contains a PAN, the certified key Ka, the PANSecret PS,
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   157
    a number specifying the target use for Ka, the signing key signK.
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   158
 *)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   159
  certC :: "[nat, key, nat, msg, key] \<Rightarrow> msg"
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   160
  where "certC PAN Ka PS T signK =
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   161
    signCert signK \<lbrace>Hash \<lbrace>Nonce PS, Pan PAN\<rbrace>, Key Ka, T\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   162
35068
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
   163
(*cert and certA have no repeated elements, so they could be abbreviations,
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
   164
  but that's tricky and makes proofs slower*)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   165
35068
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
   166
abbreviation "onlyEnc == Number 0"
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
   167
abbreviation "onlySig == Number (Suc 0)"
544867142ea4 modernized translations;
wenzelm
parents: 33028
diff changeset
   168
abbreviation "authCode == Number (Suc (Suc 0))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   169
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   170
subsection\<open>Encryption Primitives\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   171
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   172
definition EXcrypt :: "[key,key,msg,msg] \<Rightarrow> msg" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   173
  \<comment> \<open>Extra Encryption\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   174
    (*K: the symmetric key   EK: the public encryption key*)
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   175
    "EXcrypt K EK M m =
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   176
       \<lbrace>Crypt K \<lbrace>M, Hash m\<rbrace>, Crypt EK \<lbrace>Key K, m\<rbrace>\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   177
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   178
definition EXHcrypt :: "[key,key,msg,msg] \<Rightarrow> msg" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   179
  \<comment> \<open>Extra Encryption with Hashing\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   180
    (*K: the symmetric key   EK: the public encryption key*)
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   181
    "EXHcrypt K EK M m =
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   182
       \<lbrace>Crypt K \<lbrace>M, Hash m\<rbrace>, Crypt EK \<lbrace>Key K, m, Hash M\<rbrace>\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   183
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   184
definition Enc :: "[key,key,key,msg] \<Rightarrow> msg" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   185
  \<comment> \<open>Simple Encapsulation with SIGNATURE\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   186
    (*SK: the sender's signing key
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   187
      K: the symmetric key
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   188
      EK: the public encryption key*)
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   189
    "Enc SK K EK M =
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   190
       \<lbrace>Crypt K (sign SK M), Crypt EK (Key K)\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   191
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   192
definition EncB :: "[key,key,key,msg,msg] \<Rightarrow> msg" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   193
  \<comment> \<open>Encapsulation with Baggage.  Keys as above, and baggage b.\<close>
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   194
    "EncB SK K EK M b =
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   195
       \<lbrace>Enc SK K EK \<lbrace>M, Hash b\<rbrace>, b\<rbrace>"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   196
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   197
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   198
subsection\<open>Basic Properties of pubEK, pubSK, priEK and priSK\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   200
lemma publicKey_eq_iff [iff]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   201
     "(publicKey b A = publicKey b' A') = (b=b' \<and> A=A')"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   202
by (blast dest: injective_publicKey)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   203
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   204
lemma privateKey_eq_iff [iff]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   205
     "(invKey (publicKey b A) = invKey (publicKey b' A')) = (b=b' \<and> A=A')"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   206
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   207
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   208
lemma not_symKeys_publicKey [iff]: "publicKey b A \<notin> symKeys"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   209
by (simp add: symKeys_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   210
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   211
lemma not_symKeys_privateKey [iff]: "invKey (publicKey b A) \<notin> symKeys"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   212
by (simp add: symKeys_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   213
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   214
lemma symKeys_invKey_eq [simp]: "K \<in> symKeys \<Longrightarrow> invKey K = K"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   215
by (simp add: symKeys_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   216
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   217
lemma symKeys_invKey_iff [simp]: "(invKey K \<in> symKeys) = (K \<in> symKeys)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   218
by (unfold symKeys_def, auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   219
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   220
text\<open>Can be slow (or even loop) as a simprule\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   221
lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) \<Longrightarrow> K \<noteq> K'"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   222
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   223
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   224
text\<open>These alternatives to \<open>symKeys_neq_imp_neq\<close> don't seem any better
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   225
in practice.\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   226
lemma publicKey_neq_symKey: "K \<in> symKeys \<Longrightarrow> publicKey b A \<noteq> K"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   227
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   228
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   229
lemma symKey_neq_publicKey: "K \<in> symKeys \<Longrightarrow> K \<noteq> publicKey b A"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   230
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   231
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   232
lemma privateKey_neq_symKey: "K \<in> symKeys \<Longrightarrow> invKey (publicKey b A) \<noteq> K"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   233
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   234
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   235
lemma symKey_neq_privateKey: "K \<in> symKeys \<Longrightarrow> K \<noteq> invKey (publicKey b A)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   236
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   237
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   238
lemma analz_symKeys_Decrypt:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   239
     "[| Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H |]  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   240
      ==> X \<in> analz H"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   241
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   242
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   243
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   244
subsection\<open>"Image" Equations That Hold for Injective Functions\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   245
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   246
lemma invKey_image_eq [iff]: "(invKey x \<in> invKey`A) = (x\<in>A)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   247
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   248
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   249
text\<open>holds because invKey is injective\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   250
lemma publicKey_image_eq [iff]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   251
     "(publicKey b A \<in> publicKey c ` AS) = (b=c \<and> A\<in>AS)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   252
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   253
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   254
lemma privateKey_image_eq [iff]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   255
     "(invKey (publicKey b A) \<in> invKey ` publicKey c ` AS) = (b=c \<and> A\<in>AS)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   256
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   257
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   258
lemma privateKey_notin_image_publicKey [iff]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   259
     "invKey (publicKey b A) \<notin> publicKey c ` AS"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   260
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   261
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   262
lemma publicKey_notin_image_privateKey [iff]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   263
     "publicKey b A \<notin> invKey ` publicKey c ` AS"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   264
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   265
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   266
lemma keysFor_parts_initState [simp]: "keysFor (parts (initState C)) = {}"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   267
apply (simp add: keysFor_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   268
apply (induct_tac "C")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   269
apply (auto intro: range_eqI)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   270
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   271
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   272
text\<open>for proving \<open>new_keys_not_used\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   273
lemma keysFor_parts_insert:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   274
     "[| K \<in> keysFor (parts (insert X H));  X \<in> synth (analz H) |]  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   275
      ==> K \<in> keysFor (parts H) | Key (invKey K) \<in> parts H"
14218
db95d1c2f51b removal of junk and improvement of the document
paulson
parents: 14206
diff changeset
   276
by (force dest!: 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   277
         parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   278
         analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   279
            intro: analz_into_parts)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   280
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   281
lemma Crypt_imp_keysFor [intro]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   282
     "[|K \<in> symKeys; Crypt K X \<in> H|] ==> K \<in> keysFor H"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   283
by (drule Crypt_imp_invKey_keysFor, simp)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   284
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   285
text\<open>Agents see their own private keys!\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   286
lemma privateKey_in_initStateCA [iff]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   287
     "Key (invKey (publicKey b A)) \<in> initState A"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   288
by (case_tac "A", auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   289
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   290
text\<open>Agents see their own public keys!\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   291
lemma publicKey_in_initStateCA [iff]: "Key (publicKey b A) \<in> initState A"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   292
by (case_tac "A", auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   293
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   294
text\<open>RCA sees CAs' public keys!\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   295
lemma pubK_CA_in_initState_RCA [iff]:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   296
     "Key (publicKey b (CA i)) \<in> initState RCA"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   297
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   298
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   299
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   300
text\<open>Spy knows all public keys\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   301
lemma knows_Spy_pubEK_i [iff]: "Key (publicKey b A) \<in> knows Spy evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   302
apply (induct_tac "evs")
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63167
diff changeset
   303
apply (simp_all add: imageI knows_Cons split: event.split)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   304
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   305
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   306
declare knows_Spy_pubEK_i [THEN analz.Inj, iff]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   307
                            (*needed????*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   308
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   309
text\<open>Spy sees private keys of bad agents! [and obviously public keys too]\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   310
lemma knows_Spy_bad_privateKey [intro!]:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   311
     "A \<in> bad \<Longrightarrow> Key (invKey (publicKey b A)) \<in> knows Spy evs"
14206
77bf175f5145 Tidying of SET's "possibility theorems" (removal of Key_supply_ax)
paulson
parents: 14199
diff changeset
   312
by (rule initState_subset_knows [THEN subsetD], simp)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   313
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   314
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   315
subsection\<open>Fresh Nonces for Possibility Theorems\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   316
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   317
lemma Nonce_notin_initState [iff]: "Nonce N \<notin> parts (initState B)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   318
by (induct_tac "B", auto)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   319
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   320
lemma Nonce_notin_used_empty [simp]: "Nonce N \<notin> used []"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   321
by (simp add: used_Nil)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   322
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   323
text\<open>In any trace, there is an upper bound N on the greatest nonce in use.\<close>
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   324
lemma Nonce_supply_lemma: "\<exists>N. \<forall>n. N\<le>n \<longrightarrow> Nonce n \<notin> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   325
apply (induct_tac "evs")
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   326
apply (rule_tac x = 0 in exI)
63648
f9f3006a5579 "split add" -> "split"
nipkow
parents: 63167
diff changeset
   327
apply (simp_all add: used_Cons split: event.split, safe)
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   328
apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   329
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   330
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   331
lemma Nonce_supply1: "\<exists>N. Nonce N \<notin> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   332
by (rule Nonce_supply_lemma [THEN exE], blast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   333
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   334
lemma Nonce_supply: "Nonce (SOME N. Nonce N \<notin> used evs) \<notin> used evs"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   335
apply (rule Nonce_supply_lemma [THEN exE])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   336
apply (rule someI, fast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   337
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   338
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   339
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   340
subsection\<open>Specialized Methods for Possibility Theorems\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   341
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   342
ML
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   343
\<open>
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   344
(*Tactic for possibility theorems*)
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   345
fun possibility_tac ctxt =
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   346
    REPEAT (*omit used_Says so that Nonces start from different traces!*)
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45827
diff changeset
   347
    (ALLGOALS (simp_tac (ctxt delsimps [@{thm used_Says}, @{thm used_Notes}]))
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   348
     THEN
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   349
     REPEAT_FIRST (eq_assume_tac ORELSE' 
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   350
                   resolve_tac ctxt [refl, conjI, @{thm Nonce_supply}]))
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   351
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   352
(*For harder protocols (such as SET_CR!), where we have to set up some
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   353
  nonces and keys initially*)
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   354
fun basic_possibility_tac ctxt =
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   355
    REPEAT 
51717
9e7d1c139569 simplifier uses proper Proof.context instead of historic type simpset;
wenzelm
parents: 45827
diff changeset
   356
    (ALLGOALS (asm_simp_tac (ctxt setSolver safe_solver))
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   357
     THEN
59498
50b60f501b05 proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents: 58889
diff changeset
   358
     REPEAT_FIRST (resolve_tac ctxt [refl, conjI]))
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   359
\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   360
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   361
method_setup possibility = \<open>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   362
    Scan.succeed (SIMPLE_METHOD o possibility_tac)\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   363
    "for proving possibility theorems"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   364
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   365
method_setup basic_possibility = \<open>
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   366
    Scan.succeed (SIMPLE_METHOD o basic_possibility_tac)\<close>
30607
c3d1590debd8 eliminated global SIMPSET, CLASET etc. -- refer to explicit context;
wenzelm
parents: 30549
diff changeset
   367
    "for proving possibility theorems"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   368
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   369
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   370
subsection\<open>Specialized Rewriting for Theorems About \<^term>\<open>analz\<close> and Image\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   371
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   372
lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} \<union> H"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   373
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   374
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   375
lemma insert_Key_image:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   376
     "insert (Key K) (Key`KK \<union> C) = Key ` (insert K KK) \<union> C"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   377
by blast
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   378
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   379
text\<open>Needed for \<open>DK_fresh_not_KeyCryptKey\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   380
lemma publicKey_in_used [iff]: "Key (publicKey b A) \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   381
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   382
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   383
lemma privateKey_in_used [iff]: "Key (invKey (publicKey b A)) \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   384
by (blast intro!: initState_into_used)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   385
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   386
text\<open>Reverse the normal simplification of "image" to build up (not break down)
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   387
  the set of keys.  Based on \<open>analz_image_freshK_ss\<close>, but simpler.\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   388
lemmas analz_image_keys_simps =
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 63648
diff changeset
   389
       simp_thms mem_simps \<comment> \<open>these two allow its use with \<open>only:\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   390
       image_insert [THEN sym] image_Un [THEN sym] 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   391
       rangeI symKeys_neq_imp_neq
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   392
       insert_Key_singleton insert_Key_image Un_assoc [THEN sym]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   393
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   394
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   395
(*General lemmas proved by Larry*)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   396
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   397
subsection\<open>Controlled Unfolding of Abbreviations\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   398
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   399
text\<open>A set is expanded only if a relation is applied to it\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   400
lemma def_abbrev_simp_relation:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   401
     "A = B \<Longrightarrow> (A \<in> X) = (B \<in> X) \<and>  
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   402
                 (u = A) = (u = B) \<and>  
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   403
                 (A = u) = (B = u)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   404
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   405
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   406
text\<open>A set is expanded only if one of the given functions is applied to it\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   407
lemma def_abbrev_simp_function:
36866
426d5781bb25 modernized specifications;
wenzelm
parents: 35416
diff changeset
   408
     "A = B  
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   409
      \<Longrightarrow> parts (insert A X) = parts (insert B X) \<and>  
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   410
          analz (insert A X) = analz (insert B X) \<and>  
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   411
          keysFor (insert A X) = keysFor (insert B X)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   412
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   413
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   414
subsubsection\<open>Special Simplification Rules for \<^term>\<open>signCert\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   415
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   416
text\<open>Avoids duplicating X and its components!\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   417
lemma parts_insert_signCert:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   418
     "parts (insert (signCert K X) H) =  
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   419
      insert \<lbrace>X, Crypt K X\<rbrace> (parts (insert (Crypt K X) H))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   420
by (simp add: signCert_def insert_commute [of X])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   421
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   422
text\<open>Avoids a case split! [X is always available]\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   423
lemma analz_insert_signCert:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   424
     "analz (insert (signCert K X) H) =  
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   425
      insert \<lbrace>X, Crypt K X\<rbrace> (insert (Crypt K X) (analz (insert X H)))"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   426
by (simp add: signCert_def insert_commute [of X])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   427
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   428
lemma keysFor_insert_signCert: "keysFor (insert (signCert K X) H) = keysFor H"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   429
by (simp add: signCert_def)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   430
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   431
text\<open>Controlled rewrite rules for \<^term>\<open>signCert\<close>, just the definitions
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   432
  of the others. Encryption primitives are just expanded, despite their huge
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   433
  redundancy!\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   434
lemmas abbrev_simps [simp] =
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   435
    parts_insert_signCert analz_insert_signCert keysFor_insert_signCert
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   436
    sign_def     [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   437
    sign_def     [THEN def_abbrev_simp_function]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   438
    signCert_def [THEN def_abbrev_simp_relation]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   439
    signCert_def [THEN def_abbrev_simp_function]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   440
    certC_def    [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   441
    certC_def    [THEN def_abbrev_simp_function]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   442
    cert_def     [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   443
    cert_def     [THEN def_abbrev_simp_function]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   444
    EXcrypt_def  [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   445
    EXcrypt_def  [THEN def_abbrev_simp_function]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   446
    EXHcrypt_def [THEN def_abbrev_simp_relation]
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   447
    EXHcrypt_def [THEN def_abbrev_simp_function]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   448
    Enc_def      [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   449
    Enc_def      [THEN def_abbrev_simp_function]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   450
    EncB_def     [THEN def_abbrev_simp_relation]
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   451
    EncB_def     [THEN def_abbrev_simp_function]
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   452
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   453
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   454
subsubsection\<open>Elimination Rules for Controlled Rewriting\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   455
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   456
lemma Enc_partsE: 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   457
     "!!R. [|Enc SK K EK M \<in> parts H;  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   458
             [|Crypt K (sign SK M) \<in> parts H;  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   459
               Crypt EK (Key K) \<in> parts H|] ==> R|]  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   460
           ==> R"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   461
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   462
by (unfold Enc_def, blast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   463
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   464
lemma EncB_partsE: 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   465
     "!!R. [|EncB SK K EK M b \<in> parts H;  
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   466
             [|Crypt K (sign SK \<lbrace>M, Hash b\<rbrace>) \<in> parts H;  
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   467
               Crypt EK (Key K) \<in> parts H;  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   468
               b \<in> parts H|] ==> R|]  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   469
           ==> R"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   470
by (unfold EncB_def Enc_def, blast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   471
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   472
lemma EXcrypt_partsE: 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   473
     "!!R. [|EXcrypt K EK M m \<in> parts H;  
61984
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   474
             [|Crypt K \<lbrace>M, Hash m\<rbrace> \<in> parts H;  
cdea44c775fa more symbols;
wenzelm
parents: 59498
diff changeset
   475
               Crypt EK \<lbrace>Key K, m\<rbrace> \<in> parts H|] ==> R|]  
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   476
           ==> R"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   477
by (unfold EXcrypt_def, blast)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   478
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   479
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   480
subsection\<open>Lemmas to Simplify Expressions Involving \<^term>\<open>analz\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   481
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   482
lemma analz_knows_absorb:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   483
     "Key K \<in> analz (knows Spy evs)  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   484
      ==> analz (Key ` (insert K H) \<union> knows Spy evs) =  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   485
          analz (Key ` H \<union> knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   486
by (simp add: analz_insert_eq Un_upper2 [THEN analz_mono, THEN subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   487
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   488
lemma analz_knows_absorb2:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   489
     "Key K \<in> analz (knows Spy evs)  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   490
      ==> analz (Key ` (insert X (insert K H)) \<union> knows Spy evs) =  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   491
          analz (Key ` (insert X H) \<union> knows Spy evs)"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   492
apply (subst insert_commute)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   493
apply (erule analz_knows_absorb)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   494
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   495
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   496
lemma analz_insert_subset_eq:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   497
     "[|X \<in> analz (knows Spy evs);  knows Spy evs \<subseteq> H|]  
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   498
      ==> analz (insert X H) = analz H"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   499
apply (rule analz_insert_eq)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   500
apply (blast intro: analz_mono [THEN [2] rev_subsetD])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   501
done
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   502
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   503
lemmas analz_insert_simps = 
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   504
         analz_insert_subset_eq Un_upper2
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 32149
diff changeset
   505
         subset_insertI [THEN [2] subset_trans] 
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   506
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   507
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 61984
diff changeset
   508
subsection\<open>Freshness Lemmas\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   509
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   510
lemma in_parts_Says_imp_used:
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   511
     "[|Key K \<in> parts {X}; Says A B X \<in> set evs|] ==> Key K \<in> used evs"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   512
by (blast intro: parts_trans dest!: Says_imp_knows_Spy [THEN parts.Inj])
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   513
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   514
text\<open>A useful rewrite rule with \<^term>\<open>analz_image_keys_simps\<close>\<close>
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   515
lemma Crypt_notin_image_Key: "Crypt K X \<notin> Key ` KK"
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   516
by auto
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   517
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   518
lemma fresh_notin_analz_knows_Spy:
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67443
diff changeset
   519
     "Key K \<notin> used evs \<Longrightarrow> Key K \<notin> analz (knows Spy evs)"
14199
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   520
by (auto dest: analz_into_parts)
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   521
d3b8d972a488 new session HOL-SET-Protocol
paulson
parents:
diff changeset
   522
end