src/HOL/Auth/Shared.thy
author paulson
Thu, 12 Sep 1996 10:40:05 +0200
changeset 1985 84cf16192e03
parent 1967 0ff58b41c037
child 2012 1b234f1fd9c7
permissions -rw-r--r--
Tidied many proofs, using AddIffs to let equivalences take the place of separate Intr and Elim rules. Also deleted most named clasets.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Shared
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     2
    ID:         $Id$
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     5
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     6
Theory of Shared Keys (common to all symmetric-key protocols)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     7
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
     8
Server keys; initial states of agents; new nonces and keys; function "sees" 
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
     9
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    10
IS THE Notes constructor needed??
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    11
*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    12
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    13
Shared = Message + List + 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    14
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    15
consts
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    16
  shrK    :: agent => key  (*symmetric keys*)
1965
789c12ea0b30 Stronger proofs; work for Otway-Rees
paulson
parents: 1943
diff changeset
    17
  leaked  :: nat set       (*Friendly agents whose keys have leaked to Enemy*)
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    18
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1965
diff changeset
    19
constdefs     (*Enemy and compromised agents*)
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1965
diff changeset
    20
  bad     :: agent set     "bad == insert Enemy (Friend``leaked)"
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1965
diff changeset
    21
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    22
rules
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    23
  isSym_shrK "isSymKey (shrK A)"
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    24
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    25
consts  (*Initial states of agents -- parameter of the construction*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    26
  initState :: agent => msg set
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    27
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    28
primrec initState agent
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    29
        (*Server knows all keys; other agents know only their own*)
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    30
  initState_Server  "initState Server     = Key `` range shrK"
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    31
  initState_Friend  "initState (Friend i) = {Key (shrK (Friend i))}"
1967
0ff58b41c037 "bad" set simplifies statements of many theorems
paulson
parents: 1965
diff changeset
    32
  initState_Enemy   "initState Enemy      = Key``shrK``bad"
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    33
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    34
datatype  (*Messages, and components of agent stores*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    35
  event = Says agent agent msg
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    36
        | Notes agent msg
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    37
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    38
consts  
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    39
  sees1 :: [agent, event] => msg set
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    40
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    41
primrec sees1 event
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    42
           (*First agent recalls all that it says, but NOT everything
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    43
             that is sent to it; it must note such things if/when received*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    44
  sees1_Says  "sees1 A (Says A' B X)  = (if A:{A',Enemy} then {X} else {})"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    45
          (*part of A's internal state*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    46
  sees1_Notes "sees1 A (Notes A' X)   = (if A=A' then {X} else {})"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    47
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    48
consts  
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    49
  sees :: [agent, event list] => msg set
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    50
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    51
primrec sees list
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    52
        (*Initial knowledge includes all public keys and own private key*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    53
  sees_Nil  "sees A []       = initState A"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    54
  sees_Cons "sees A (ev#evs) = sees1 A ev Un sees A evs"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    55
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    56
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    57
(*Agents generate "random" nonces.  Different traces always yield
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    58
  different nonces.  Same applies for keys.*)
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    59
consts
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    60
  newN :: "event list => nat"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    61
  newK :: "event list => key"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    62
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    63
rules
1943
20574dca5a3e Renaming and simplification
paulson
parents: 1934
diff changeset
    64
  inj_shrK "inj shrK"
1934
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    65
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    66
  inj_newN   "inj newN"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    67
  fresh_newN "Nonce (newN evs) ~: parts (initState B)" 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    68
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    69
  inj_newK   "inj newK"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    70
  fresh_newK "Key (newK evs) ~: parts (initState B)" 
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    71
  isSym_newK "isSymKey (newK evs)"
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    72
58573e7041b4 Separation of theory Event into two parts:
paulson
parents:
diff changeset
    73
end