src/HOL/Auth/Public.ML
author paulson
Thu, 11 Sep 1997 12:22:31 +0200
changeset 3667 42a726e008ce
parent 3544 6ae62d55a620
child 3673 3b06747c3f37
permissions -rw-r--r--
Now uses the generic induct_tac
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Public
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     2
    ID:         $Id$
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     5
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     6
Theory of Public Keys (common to all symmetric-key protocols)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     7
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     8
Server keys; initial states of agents; new nonces and keys; function "sees" 
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
     9
*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    10
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    11
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    12
open Public;
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    13
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    14
(*** Basic properties of pubK & priK ***)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    15
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    16
AddIffs [inj_pubK RS inj_eq];
3477
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    17
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    18
goal thy "!!A B. (priK A = priK B) = (A=B)";
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    19
by (Step_tac 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    20
by (dres_inst_tac [("f","invKey")] arg_cong 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    21
by (Full_simp_tac 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    22
qed "priK_inj_eq";
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    23
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    24
AddIffs [priK_inj_eq];
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    25
AddIffs [priK_neq_pubK, priK_neq_pubK RS not_sym];
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    26
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    27
goalw thy [isSymKey_def] "~ isSymKey (pubK A)";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    28
by (Simp_tac 1);
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    29
qed "not_isSymKey_pubK";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    30
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    31
goalw thy [isSymKey_def] "~ isSymKey (priK A)";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    32
by (Simp_tac 1);
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    33
qed "not_isSymKey_priK";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    34
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    35
AddIffs [not_isSymKey_pubK, not_isSymKey_priK];
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    36
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    37
(** Rewrites should not refer to  initState(Friend i) 
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    38
    -- not in normal form! **)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    39
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    40
goalw thy [keysFor_def] "keysFor (parts (initState C)) = {}";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    41
by (induct_tac "C" 1);
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    42
by (auto_tac (!claset addIs [range_eqI], !simpset));
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    43
qed "keysFor_parts_initState";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    44
Addsimps [keysFor_parts_initState];
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    45
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    46
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    47
(*** Function "sees" ***)
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    48
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    49
(*Agents see their own private keys!*)
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    50
goal thy "A ~= Spy --> Key (priK A) : sees A evs";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    51
by (induct_tac "evs" 1);
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    52
by (induct_tac "A" 1);
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    53
by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    54
qed_spec_mp "sees_own_priK";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    55
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    56
(*All public keys are visible to all*)
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    57
goal thy "Key (pubK A) : sees B evs";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    58
by (induct_tac "evs" 1);
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    59
by (induct_tac "B" 1);
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    60
by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    61
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    62
qed_spec_mp "sees_pubK";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    63
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    64
(*Spy sees private keys of agents!*)
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    65
goal thy "!!A. A: lost ==> Key (priK A) : sees Spy evs";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    66
by (induct_tac "evs" 1);
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    67
by (ALLGOALS (asm_simp_tac (!simpset addsimps [sees_Cons])));
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
    68
by (Blast_tac 1);
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    69
qed "Spy_sees_lost";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    70
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    71
AddIffs [sees_pubK, sees_pubK RS analz.Inj];
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    72
AddSIs  [sees_own_priK, Spy_sees_lost];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    73
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    74
3442
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    75
(*For not_lost_tac*)
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    76
goal thy "!!A. [| Crypt (pubK A) X : analz (sees Spy evs);  A: lost |] \
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    77
\              ==> X : analz (sees Spy evs)";
3442
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    78
by (blast_tac (!claset addSDs [analz.Decrypt]) 1);
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    79
qed "Crypt_Spy_analz_lost";
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    80
3442
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    81
(*Prove that the agent is uncompromised by the confidentiality of 
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    82
  a component of a message she's said.*)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    83
fun not_lost_tac s =
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    84
    case_tac ("(" ^ s ^ ") : lost") THEN'
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    85
    SELECT_GOAL 
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    86
      (REPEAT_DETERM (dtac (Says_imp_sees_Spy RS analz.Inj) 1) THEN
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    87
       REPEAT_DETERM (etac MPair_analz 1) THEN
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    88
       THEN_BEST_FIRST 
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    89
         (dres_inst_tac [("A", s)] Crypt_Spy_analz_lost 1 THEN assume_tac 1)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    90
         (has_fewer_prems 1, size_of_thm)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
    91
         (Step_tac 1));
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    92
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    93
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    94
(*** Fresh nonces ***)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    95
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    96
goal thy "Nonce N ~: parts (initState B)";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
    97
by (induct_tac "B" 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    98
by (Auto_tac ());
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    99
qed "Nonce_notin_initState";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   100
AddIffs [Nonce_notin_initState];
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   101
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   102
goal thy "Nonce N ~: used []";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   103
by (simp_tac (!simpset addsimps [used_def]) 1);
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   104
qed "Nonce_notin_used_empty";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   105
Addsimps [Nonce_notin_used_empty];
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   106
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   107
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   108
(*** Supply fresh nonces for possibility theorems. ***)
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   109
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   110
goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs";
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
   111
by (induct_tac "evs" 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   112
by (res_inst_tac [("x","0")] exI 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   113
by (Step_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   114
by (Full_simp_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   115
(*Inductive step*)
3667
42a726e008ce Now uses the generic induct_tac
paulson
parents: 3544
diff changeset
   116
by (induct_tac "a" 1);
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   117
by (ALLGOALS (full_simp_tac 
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   118
	      (!simpset delsimps [sees_Notes]
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   119
		        addsimps [UN_parts_sees_Says,
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   120
				  UN_parts_sees_Notes])));
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   121
by (ALLGOALS (rtac (msg_Nonce_supply RS exE)));
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   122
by (ALLGOALS (blast_tac (!claset addSEs [add_leE])));
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   123
val lemma = result();
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   124
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   125
goal thy "Nonce (@ N. Nonce N ~: used evs) ~: used evs";
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   126
by (rtac (lemma RS exE) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   127
by (rtac selectI 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   128
by (Fast_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   129
qed "Nonce_supply";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   130
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   131
(*Tactic for possibility theorems*)
3544
6ae62d55a620 Now possibility_tac is an explicit function, in order to delay
paulson
parents: 3519
diff changeset
   132
fun possibility_tac st = st |>
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   133
    REPEAT (*omit used_Says so that Nonces, Keys start from different traces!*)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   134
    (ALLGOALS (simp_tac (!simpset delsimps [used_Says] setSolver safe_solver))
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   135
     THEN
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   136
     REPEAT_FIRST (eq_assume_tac ORELSE' 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   137
                   resolve_tac [refl, conjI, Nonce_supply]));
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   138
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   139
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   140
(*** Specialized rewriting for the analz_image_... theorems ***)
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   141
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   142
goal thy "insert (Key K) H = Key `` {K} Un H";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   143
by (Blast_tac 1);
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   144
qed "insert_Key_singleton";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   145
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   146
goal thy "insert (Key K) (Key``KK Un C) = Key `` (insert K KK) Un C";
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   147
by (Blast_tac 1);
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   148
qed "insert_Key_image";
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   149
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   150
(*Reverse the normal simplification of "image" to build up (not break down)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   151
  the set of keys.  Based on analz_image_freshK_ss, but simpler.*)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   152
val analz_image_keys_ss = 
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   153
     !simpset delsimps [image_insert, image_Un]
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   154
              addsimps [image_insert RS sym, image_Un RS sym,
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   155
			rangeI, 
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   156
			insert_Key_singleton, 
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   157
			insert_Key_image, Un_assoc RS sym]
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents: 3477
diff changeset
   158
              setloop split_tac [expand_if];
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   159