doc-src/TutorialI/Protocol/Public.thy
author wenzelm
Fri, 13 Mar 2009 19:53:09 +0100
changeset 30514 e19d5b459a61
parent 26019 ecbfe2645694
child 30548 2eef5e71edd6
permissions -rw-r--r--
more regular method setup via SIMPLE_METHOD;
     1 (*  Title:      HOL/Auth/Public
     2     ID:         $Id$
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
     4     Copyright   1996  University of Cambridge
     5 
     6 Theory of Public Keys (common to all public-key protocols)
     7 
     8 Private and public keys; initial states of agents
     9 *)(*<*)
    10 theory Public imports Event
    11 begin
    12 (*>*)
    13 
    14 text {*
    15 The function
    16 @{text pubK} maps agents to their public keys.  The function
    17 @{text priK} maps agents to their private keys.  It is merely
    18 an abbreviation (cf.\ \S\ref{sec:abbreviations}) defined in terms of
    19 @{text invKey} and @{text pubK}.
    20 *}
    21 
    22 consts pubK :: "agent \<Rightarrow> key"
    23 abbreviation priK :: "agent \<Rightarrow> key"
    24 where "priK x  \<equiv>  invKey(pubK x)"
    25 (*<*)
    26 primrec
    27         (*Agents know their private key and all public keys*)
    28   initState_Server:  "initState Server     =    
    29  		         insert (Key (priK Server)) (Key ` range pubK)"
    30   initState_Friend:  "initState (Friend i) =    
    31  		         insert (Key (priK (Friend i))) (Key ` range pubK)"
    32   initState_Spy:     "initState Spy        =    
    33  		         (Key`invKey`pubK`bad) Un (Key ` range pubK)"
    34 (*>*)
    35 
    36 text {*
    37 \noindent
    38 The set @{text bad} consists of those agents whose private keys are known to
    39 the spy.
    40 
    41 Two axioms are asserted about the public-key cryptosystem. 
    42 No two agents have the same public key, and no private key equals
    43 any public key.
    44 *}
    45 
    46 axioms
    47   inj_pubK:        "inj pubK"
    48   priK_neq_pubK:   "priK A \<noteq> pubK B"
    49 (*<*)
    50 lemmas [iff] = inj_pubK [THEN inj_eq]
    51 
    52 lemma priK_inj_eq[iff]: "(priK A = priK B) = (A=B)"
    53   apply safe
    54   apply (drule_tac f=invKey in arg_cong)
    55   apply simp
    56   done
    57 
    58 lemmas [iff] = priK_neq_pubK priK_neq_pubK [THEN not_sym]
    59 
    60 lemma not_symKeys_pubK[iff]: "pubK A \<notin> symKeys"
    61   by (simp add: symKeys_def)
    62 
    63 lemma not_symKeys_priK[iff]: "priK A \<notin> symKeys"
    64   by (simp add: symKeys_def)
    65 
    66 lemma symKeys_neq_imp_neq: "(K \<in> symKeys) \<noteq> (K' \<in> symKeys) \<Longrightarrow> K \<noteq> K'"
    67   by blast
    68 
    69 lemma analz_symKeys_Decrypt: "[| Crypt K X \<in> analz H;  K \<in> symKeys;  Key K \<in> analz H |]
    70      ==> X \<in> analz H"
    71   by (auto simp add: symKeys_def)
    72 
    73 
    74 (** "Image" equations that hold for injective functions **)
    75 
    76 lemma invKey_image_eq[simp]: "(invKey x : invKey`A) = (x:A)"
    77   by auto
    78 
    79 (*holds because invKey is injective*)
    80 lemma pubK_image_eq[simp]: "(pubK x : pubK`A) = (x:A)"
    81   by auto
    82 
    83 lemma priK_pubK_image_eq[simp]: "(priK x ~: pubK`A)"
    84   by auto
    85 
    86 
    87 (** Rewrites should not refer to  initState(Friend i) 
    88     -- not in normal form! **)
    89 
    90 lemma keysFor_parts_initState[simp]: "keysFor (parts (initState C)) = {}"
    91   apply (unfold keysFor_def)
    92   apply (induct C)
    93   apply (auto intro: range_eqI)
    94   done
    95 
    96 
    97 (*** Function "spies" ***)
    98 
    99 (*Agents see their own private keys!*)
   100 lemma priK_in_initState[iff]: "Key (priK A) : initState A"
   101   by (induct A) auto
   102 
   103 (*All public keys are visible*)
   104 lemma spies_pubK[iff]: "Key (pubK A) : spies evs"
   105   by (induct evs) (simp_all add: imageI knows_Cons split: event.split)
   106 
   107 (*Spy sees private keys of bad agents!*)
   108 lemma Spy_spies_bad[intro!]: "A: bad ==> Key (priK A) : spies evs"
   109   by (induct evs) (simp_all add: imageI knows_Cons split: event.split)
   110 
   111 lemmas [iff] = spies_pubK [THEN analz.Inj]
   112 
   113 
   114 (*** Fresh nonces ***)
   115 
   116 lemma Nonce_notin_initState[iff]: "Nonce N ~: parts (initState B)"
   117   by (induct B) auto
   118 
   119 lemma Nonce_notin_used_empty[simp]: "Nonce N ~: used []"
   120   by (simp add: used_Nil)
   121 
   122 
   123 (*** Supply fresh nonces for possibility theorems. ***)
   124 
   125 (*In any trace, there is an upper bound N on the greatest nonce in use.*)
   126 lemma Nonce_supply_lemma: "EX N. ALL n. N<=n --> Nonce n \<notin> used evs"
   127 apply (induct_tac "evs")
   128 apply (rule_tac x = 0 in exI)
   129 apply (simp_all (no_asm_simp) add: used_Cons split add: event.split)
   130 apply safe
   131 apply (rule msg_Nonce_supply [THEN exE], blast elim!: add_leE)+
   132 done
   133 
   134 lemma Nonce_supply1: "EX N. Nonce N \<notin> used evs"
   135 by (rule Nonce_supply_lemma [THEN exE], blast)
   136 
   137 lemma Nonce_supply: "Nonce (@ N. Nonce N \<notin> used evs) \<notin> used evs"
   138 apply (rule Nonce_supply_lemma [THEN exE])
   139 apply (rule someI, fast)
   140 done
   141 
   142 
   143 (*** Specialized rewriting for the analz_image_... theorems ***)
   144 
   145 lemma insert_Key_singleton: "insert (Key K) H = Key ` {K} Un H"
   146   by blast
   147 
   148 lemma insert_Key_image: "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C"
   149   by blast
   150 
   151 
   152 (*Specialized methods*)
   153 
   154 (*Tactic for possibility theorems*)
   155 ML {*
   156 fun possibility_tac st = st |>
   157     REPEAT (*omit used_Says so that Nonces start from different traces!*)
   158     (ALLGOALS (simp_tac (@{simpset} delsimps [used_Says]))
   159      THEN
   160      REPEAT_FIRST (eq_assume_tac ORELSE' 
   161                    resolve_tac [refl, conjI, @{thm Nonce_supply}]));
   162 *}
   163 
   164 method_setup possibility = {* Method.no_args (SIMPLE_METHOD possibility_tac) *}
   165     "for proving possibility theorems"
   166 
   167 end
   168 (*>*)