src/HOL/Auth/Public_lemmas.ML
changeset 13922 75ae4244a596
parent 13921 69c627b6b28d
child 13923 019342d03d81
equal deleted inserted replaced
13921:69c627b6b28d 13922:75ae4244a596
     1 (*  Title:      HOL/Auth/Public_lemmas
       
     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 symmetric-key protocols)
       
     7 
       
     8 Server keys; initial states of agents; new nonces and keys; function "sees" 
       
     9 *)
       
    10 
       
    11 val inj_pubK      = thm "inj_pubK";
       
    12 val priK_neq_pubK = thm "priK_neq_pubK";
       
    13 
       
    14 (*** Basic properties of pubK & priK ***)
       
    15 
       
    16 AddIffs [inj_pubK RS inj_eq];
       
    17 
       
    18 Goal "(priK A = priK B) = (A=B)";
       
    19 by Safe_tac;
       
    20 by (dres_inst_tac [("f","invKey")] arg_cong 1);
       
    21 by (Full_simp_tac 1);
       
    22 qed "priK_inj_eq";
       
    23 
       
    24 AddIffs [priK_inj_eq];
       
    25 AddIffs [priK_neq_pubK, priK_neq_pubK RS not_sym];
       
    26 
       
    27 Goalw [symKeys_def] "pubK A \\<notin> symKeys";
       
    28 by (Simp_tac 1);
       
    29 qed "not_symKeys_pubK";
       
    30 
       
    31 Goalw [symKeys_def] "priK A \\<notin> symKeys";
       
    32 by (Simp_tac 1);
       
    33 qed "not_symKeys_priK";
       
    34 
       
    35 AddIffs [not_symKeys_pubK, not_symKeys_priK];
       
    36 
       
    37 Goal "(K \\<in> symKeys) \\<noteq> (K' \\<in> symKeys) ==> K \\<noteq> K'";
       
    38 by (Blast_tac 1); 
       
    39 qed "symKeys_neq_imp_neq";
       
    40 
       
    41 Goal "[| Crypt K X \\<in> analz H;  K \\<in> symKeys;  Key K \\<in> analz H |] \
       
    42 \     ==> X \\<in> analz H";
       
    43 by (auto_tac(claset(), simpset() addsimps [symKeys_def]));
       
    44 qed "analz_symKeys_Decrypt";
       
    45 
       
    46 
       
    47 (** "Image" equations that hold for injective functions **)
       
    48 
       
    49 Goal "(invKey x : invKey`A) = (x:A)";
       
    50 by Auto_tac;
       
    51 qed "invKey_image_eq";
       
    52 
       
    53 (*holds because invKey is injective*)
       
    54 Goal "(pubK x : pubK`A) = (x:A)";
       
    55 by Auto_tac;
       
    56 qed "pubK_image_eq";
       
    57 
       
    58 Goal "(priK x ~: pubK`A)";
       
    59 by Auto_tac;
       
    60 qed "priK_pubK_image_eq";
       
    61 Addsimps [invKey_image_eq, pubK_image_eq, priK_pubK_image_eq];
       
    62 
       
    63 
       
    64 (** Rewrites should not refer to  initState(Friend i) 
       
    65     -- not in normal form! **)
       
    66 
       
    67 Goalw [keysFor_def] "keysFor (parts (initState C)) = {}";
       
    68 by (induct_tac "C" 1);
       
    69 by (auto_tac (claset() addIs [range_eqI], simpset()));
       
    70 qed "keysFor_parts_initState";
       
    71 Addsimps [keysFor_parts_initState];
       
    72 
       
    73 
       
    74 (*** Function "spies" ***)
       
    75 
       
    76 (*Agents see their own private keys!*)
       
    77 Goal "Key (priK A) : initState A";
       
    78 by (induct_tac "A" 1);
       
    79 by Auto_tac;
       
    80 qed "priK_in_initState";
       
    81 AddIffs [priK_in_initState];
       
    82 
       
    83 (*All public keys are visible*)
       
    84 Goal "Key (pubK A) : spies evs";
       
    85 by (induct_tac "evs" 1);
       
    86 by (ALLGOALS (asm_simp_tac
       
    87 	      (simpset() addsimps [imageI, knows_Cons]
       
    88 	                addsplits [expand_event_case])));
       
    89 qed_spec_mp "spies_pubK";
       
    90 
       
    91 (*Spy sees private keys of bad agents!*)
       
    92 Goal "A: bad ==> Key (priK A) : spies evs";
       
    93 by (induct_tac "evs" 1);
       
    94 by (ALLGOALS (asm_simp_tac
       
    95 	      (simpset() addsimps [imageI, knows_Cons]
       
    96 	                addsplits [expand_event_case])));
       
    97 qed "Spy_spies_bad";
       
    98 
       
    99 AddIffs [spies_pubK, spies_pubK RS analz.Inj];
       
   100 AddSIs  [Spy_spies_bad];
       
   101 
       
   102 
       
   103 (*** Fresh nonces ***)
       
   104 
       
   105 Goal "Nonce N ~: parts (initState B)";
       
   106 by (induct_tac "B" 1);
       
   107 by Auto_tac;
       
   108 qed "Nonce_notin_initState";
       
   109 AddIffs [Nonce_notin_initState];
       
   110 
       
   111 Goal "Nonce N ~: used []";
       
   112 by (simp_tac (simpset() addsimps [used_Nil]) 1);
       
   113 qed "Nonce_notin_used_empty";
       
   114 Addsimps [Nonce_notin_used_empty];
       
   115 
       
   116 
       
   117 (*** Supply fresh nonces for possibility theorems. ***)
       
   118 
       
   119 (*In any trace, there is an upper bound N on the greatest nonce in use.*)
       
   120 Goal "EX N. ALL n. N<=n --> Nonce n ~: used evs";
       
   121 by (induct_tac "evs" 1);
       
   122 by (res_inst_tac [("x","0")] exI 1);
       
   123 by (ALLGOALS (asm_simp_tac
       
   124 	      (simpset() addsimps [used_Cons]
       
   125 			addsplits [expand_event_case])));
       
   126 by Safe_tac;
       
   127 by (ALLGOALS (rtac (msg_Nonce_supply RS exE)));
       
   128 by (ALLGOALS (blast_tac (claset() addSEs [add_leE])));
       
   129 val lemma = result();
       
   130 
       
   131 Goal "EX N. Nonce N ~: used evs";
       
   132 by (rtac (lemma RS exE) 1);
       
   133 by (Blast_tac 1);
       
   134 qed "Nonce_supply1";
       
   135 
       
   136 Goal "Nonce (@ N. Nonce N ~: used evs) ~: used evs";
       
   137 by (rtac (lemma RS exE) 1);
       
   138 by (rtac someI 1);
       
   139 by (Fast_tac 1);
       
   140 qed "Nonce_supply";
       
   141 
       
   142 (*Tactic for possibility theorems (Isar interface)*)
       
   143 fun gen_possibility_tac ss state = state |>
       
   144     REPEAT (*omit used_Says so that Nonces start from different traces!*)
       
   145     (ALLGOALS (simp_tac (ss delsimps [used_Says]))
       
   146      THEN
       
   147      REPEAT_FIRST (eq_assume_tac ORELSE' 
       
   148                    resolve_tac [refl, conjI, Nonce_supply]));
       
   149 
       
   150 (*Tactic for possibility theorems (ML script version)*)
       
   151 fun possibility_tac state = gen_possibility_tac (simpset()) state;
       
   152 
       
   153 (*** Specialized rewriting for the analz_image_... theorems ***)
       
   154 
       
   155 Goal "insert (Key K) H = Key ` {K} Un H";
       
   156 by (Blast_tac 1);
       
   157 qed "insert_Key_singleton";
       
   158 
       
   159 Goal "insert (Key K) (Key`KK Un C) = Key ` (insert K KK) Un C";
       
   160 by (Blast_tac 1);
       
   161 qed "insert_Key_image";
       
   162 
       
   163 (*Reverse the normal simplification of "image" to build up (not break down)
       
   164   the set of keys.  Based on analz_image_freshK_ss, but simpler.*)
       
   165 val analz_image_keys_ss = 
       
   166      simpset() delsimps [image_insert, image_Un]
       
   167 	       delsimps [imp_disjL]    (*reduces blow-up*)
       
   168 	       addsimps [image_insert RS sym, image_Un RS sym,
       
   169 			 rangeI, 
       
   170 			 insert_Key_singleton, 
       
   171 			 insert_Key_image, Un_assoc RS sym];
       
   172