src/HOL/Auth/Public.ML
author paulson
Tue, 01 Jul 1997 17:34:42 +0200
changeset 3477 3aced7fa7d8b
parent 3473 c2334f9532ab
child 3512 9dcb4daa15e8
permissions -rw-r--r--
New theorem priK_inj_eq, injectivity of priK
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
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    14
(*Holds because Friend is injective: thus cannot prove for all f*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    15
goal thy "(Friend x : Friend``A) = (x:A)";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    16
by (Auto_tac());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    17
qed "Friend_image_eq";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    18
Addsimps [Friend_image_eq];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    19
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    20
Addsimps [Un_insert_left, Un_insert_right];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    21
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    22
(*By default only o_apply is built-in.  But in the presence of eta-expansion
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    23
  this means that some terms displayed as (f o g) will be rewritten, and others
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    24
  will not!*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    25
Addsimps [o_def];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    26
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    27
goalw thy [keysFor_def] "keysFor (parts (initState lost C)) = {}";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    28
by (agent.induct_tac "C" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    29
by (auto_tac (!claset addIs [range_eqI], !simpset));
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    30
qed "keysFor_parts_initState";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    31
Addsimps [keysFor_parts_initState];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    32
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    33
goalw thy [keysFor_def] "keysFor (Key``E) = {}";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    34
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    35
qed "keysFor_image_Key";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    36
Addsimps [keysFor_image_Key];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    37
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    38
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    39
(*** Function "sees" ***)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    40
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    41
goal thy
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    42
    "!!evs. lost' <= lost ==> sees lost' A evs <= sees lost A evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    43
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    44
by (agent.induct_tac "A" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    45
by (event.induct_tac "a" 2);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    46
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    47
qed "sees_mono";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    48
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    49
(*** Basic properties of pubK & priK ***)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    50
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    51
AddIffs [inj_pubK RS inj_eq];
3477
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    52
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    53
goal thy "!!A B. (priK A = priK B) = (A=B)";
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    54
by (Step_tac 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    55
by (dres_inst_tac [("f","invKey")] arg_cong 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    56
by (Full_simp_tac 1);
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    57
qed "priK_inj_eq";
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    58
3aced7fa7d8b New theorem priK_inj_eq, injectivity of priK
paulson
parents: 3473
diff changeset
    59
AddIffs [priK_inj_eq];
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    60
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
    61
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    62
goalw thy [isSymKey_def] "~ isSymKey (pubK A)";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    63
by (Simp_tac 1);
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    64
qed "not_isSymKey_pubK";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    65
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    66
goalw thy [isSymKey_def] "~ isSymKey (priK A)";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    67
by (Simp_tac 1);
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    68
qed "not_isSymKey_priK";
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    69
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    70
AddIffs [not_isSymKey_pubK, not_isSymKey_priK];
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    71
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
    72
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    73
(*Agents see their own private keys!*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    74
goal thy "A ~= Spy --> Key (priK A) : sees lost A evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    75
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    76
by (agent.induct_tac "A" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    77
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    78
qed_spec_mp "sees_own_priK";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    79
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    80
(*All public keys are visible to all*)
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    81
goal thy "Key (pubK A) : sees lost B evs";
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    82
by (list.induct_tac "evs" 1);
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    83
by (agent.induct_tac "B" 1);
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    84
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    85
qed_spec_mp "sees_pubK";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    86
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    87
(*Spy sees private keys of lost agents!*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    88
goal thy "!!A. A: lost ==> Key (priK A) : sees lost Spy evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    89
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    90
by (Auto_tac());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    91
qed "Spy_sees_lost";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    92
3473
c2334f9532ab New and stronger lemmas; more default simp/cla rules
paulson
parents: 3465
diff changeset
    93
AddIffs [sees_pubK, sees_pubK RS analz.Inj];
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    94
AddSIs  [sees_own_priK, Spy_sees_lost];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    95
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    96
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    97
(** Specialized rewrite rules for (sees lost A (Says...#evs)) **)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    98
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
    99
goal thy "sees lost B (Says A B X # evs) = insert X (sees lost B evs)";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   100
by (Simp_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   101
qed "sees_own";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   102
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   103
goal thy "!!A. Server ~= B ==> \
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   104
\          sees lost Server (Says A B X # evs) = sees lost Server evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   105
by (Asm_simp_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   106
qed "sees_Server";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   107
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   108
goal thy "!!A. Friend i ~= B ==> \
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   109
\          sees lost (Friend i) (Says A B X # evs) = sees lost (Friend i) evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   110
by (Asm_simp_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   111
qed "sees_Friend";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   112
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   113
goal thy "sees lost Spy (Says A B X # evs) = insert X (sees lost Spy evs)";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   114
by (Simp_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   115
qed "sees_Spy";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   116
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   117
goal thy "sees lost A (Says A' B X # evs) <= insert X (sees lost A evs)";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   118
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   119
by (Fast_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   120
qed "sees_Says_subset_insert";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   121
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   122
goal thy "sees lost A evs <= sees lost A (Says A' B X # evs)";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   123
by (simp_tac (!simpset setloop split_tac [expand_if]) 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   124
by (Fast_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   125
qed "sees_subset_sees_Says";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   126
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   127
(*Pushing Unions into parts.  One of the agents A is B, and thus sees Y.
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   128
  Once used to prove new_keys_not_seen; now obsolete.*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   129
goal thy "(UN A. parts (sees lost A (Says B C Y # evs))) = \
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   130
\         parts {Y} Un (UN A. parts (sees lost A evs))";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   131
by (Step_tac 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   132
by (etac rev_mp 1);     (*split_tac does not work on assumptions*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   133
by (ALLGOALS
3207
fe79ad367d77 renamed unsafe_addss to addss
oheimb
parents: 3121
diff changeset
   134
    (fast_tac (!claset addss (!simpset addsimps [parts_Un, sees_Cons] 
3121
cbb6c0c1c58a Conversion to use blast_tac (with other improvements)
paulson
parents: 2637
diff changeset
   135
				            setloop split_tac [expand_if]))));
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   136
qed "UN_parts_sees_Says";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   137
3465
e85c24717cad set_of_list -> set
nipkow
parents: 3442
diff changeset
   138
goal thy "Says A B X : set evs --> X : sees lost Spy evs";
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   139
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   140
by (Auto_tac ());
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   141
qed_spec_mp "Says_imp_sees_Spy";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   142
2537
906704a5b3bf Added sees_Spy_partsEs
paulson
parents: 2516
diff changeset
   143
(*Use with addSEs to derive contradictions from old Says events containing
906704a5b3bf Added sees_Spy_partsEs
paulson
parents: 2516
diff changeset
   144
  items known to be fresh*)
906704a5b3bf Added sees_Spy_partsEs
paulson
parents: 2516
diff changeset
   145
val sees_Spy_partsEs = make_elim (Says_imp_sees_Spy RS parts.Inj):: partsEs;
906704a5b3bf Added sees_Spy_partsEs
paulson
parents: 2516
diff changeset
   146
3442
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   147
(*For not_lost_tac*)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   148
goal thy "!!A. [| Crypt (pubK A) X : analz (sees lost Spy evs);  A: lost |] \
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   149
\              ==> X : analz (sees lost Spy evs)";
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   150
by (blast_tac (!claset addSDs [analz.Decrypt]) 1);
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   151
qed "Crypt_Spy_analz_lost";
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   152
3442
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   153
(*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
   154
  a component of a message she's said.*)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   155
fun not_lost_tac s =
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   156
    case_tac ("(" ^ s ^ ") : lost") THEN'
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   157
    SELECT_GOAL 
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   158
      (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
   159
       REPEAT_DETERM (etac MPair_analz 1) THEN
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   160
       THEN_BEST_FIRST 
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   161
         (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
   162
         (has_fewer_prems 1, size_of_thm)
0b804b390b0e Removed Says_Crypt_lost and Says_Crypt_not_lost.
paulson
parents: 3207
diff changeset
   163
         (Step_tac 1));
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   164
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   165
Addsimps [sees_own, sees_Server, sees_Friend, sees_Spy];
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   166
Delsimps [sees_Cons];   (**** NOTE REMOVAL -- laws above are cleaner ****)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   167
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   168
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   169
(*** Fresh nonces ***)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   170
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   171
goal thy "Nonce N ~: parts (initState lost B)";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   172
by (agent.induct_tac "B" 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   173
by (Auto_tac ());
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   174
qed "Nonce_notin_initState";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   175
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   176
AddIffs [Nonce_notin_initState];
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   177
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   178
goalw thy [used_def] "!!X. X: parts (sees lost B evs) ==> X: used evs";
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   179
by (etac (impOfSubs parts_mono) 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   180
by (Fast_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   181
qed "usedI";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   182
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   183
AddIs [usedI];
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   184
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   185
(** A supply of fresh nonces for possibility theorems. **)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   186
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   187
goalw thy [used_def] "EX N. ALL n. N<=n --> Nonce n ~: used evs";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   188
by (list.induct_tac "evs" 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   189
by (res_inst_tac [("x","0")] exI 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   190
by (Step_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   191
by (Full_simp_tac 1);
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   192
(*Inductive step*)
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   193
by (event.induct_tac "a" 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   194
by (full_simp_tac (!simpset addsimps [UN_parts_sees_Says]) 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   195
by (msg.induct_tac "msg" 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   196
by (ALLGOALS (asm_simp_tac (!simpset addsimps [exI, parts_insert2])));
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   197
by (Step_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   198
(*MPair case*)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   199
by (res_inst_tac [("x","Na+Nb")] exI 2);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   200
by (fast_tac (!claset addSEs [add_leE]) 2);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   201
(*Nonce case*)
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   202
by (res_inst_tac [("x","N + Suc nat")] exI 1);
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2537
diff changeset
   203
by (fast_tac (!claset addSEs [add_leE] addaltern trans_tac) 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   204
val lemma = result();
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   205
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   206
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
   207
by (rtac (lemma RS exE) 1);
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   208
by (rtac selectI 1);
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   209
by (Fast_tac 1);
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   210
qed "Nonce_supply";
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   211
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   212
(*Tactic for possibility theorems*)
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   213
val possibility_tac =
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   214
    REPEAT 
2637
e9b203f854ae reflecting my recent changes of the simplifier and classical reasoner
oheimb
parents: 2537
diff changeset
   215
    (ALLGOALS (simp_tac (!simpset setSolver safe_solver))
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   216
     THEN
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   217
     REPEAT_FIRST (eq_assume_tac ORELSE' 
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2497
diff changeset
   218
                   resolve_tac [refl, conjI, Nonce_supply]));
2497
47de509bdd55 New treatment of nonce creation
paulson
parents: 2451
diff changeset
   219
2318
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   220
(** Power of the Spy **)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   221
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   222
(*The Spy can see more than anybody else, except for their initial state*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   223
goal thy "sees lost A evs <= initState lost A Un sees lost Spy evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   224
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   225
by (event.induct_tac "a" 2);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   226
by (ALLGOALS (fast_tac (!claset addDs [sees_Says_subset_insert RS subsetD] 
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   227
                                addss (!simpset))));
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   228
qed "sees_agent_subset_sees_Spy";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   229
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   230
(*The Spy can see more than anybody else who's lost their key!*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   231
goal thy "A: lost --> A ~= Server --> sees lost A evs <= sees lost Spy evs";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   232
by (list.induct_tac "evs" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   233
by (event.induct_tac "a" 2);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   234
by (agent.induct_tac "A" 1);
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   235
by (auto_tac (!claset addDs [sees_Says_subset_insert RS subsetD], (!simpset)));
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   236
qed_spec_mp "sees_lost_agent_subset_sees_Spy";
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   237
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   238
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   239
(** Simplifying   parts (insert X (sees lost A evs))
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   240
      = parts {X} Un parts (sees lost A evs) -- since general case loops*)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   241
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   242
val parts_insert_sees = 
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   243
    parts_insert |> read_instantiate_sg (sign_of thy)
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   244
                                        [("H", "sees lost A evs")]
6d3f7c7f70b0 Public-key examples
paulson
parents:
diff changeset
   245
                 |> standard;