src/HOL/UNITY/NSP_Bad.thy
author paulson
Thu, 03 Dec 1998 10:45:06 +0100
changeset 6012 1894bfc4aee9
parent 5596 b29d18d8c4d2
child 6216 05d99c0bbfa0
permissions -rw-r--r--
Addition of the States component; parts of Comp not working
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
5430
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/NSP_Bad
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     2
    ID:         $Id$
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     5
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     6
loadpath := "../Auth" :: !loadpath; use_thy"NSP_Bad";
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     7
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     8
Security protocols in UNITY: Needham-Schroeder, public keys (flawed version).
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
     9
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    10
Original file is ../Auth/NS_Public_Bad
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    11
*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    12
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    13
NSP_Bad = Public + Constrains + 
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    14
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    15
types state = event list
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    16
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    17
constdefs
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    18
  
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    19
  (*The spy MAY say anything he CAN say.  We do not expect him to
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    20
    invent new nonces here, but he can also use NS1.  Common to
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    21
    all similar protocols.*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    22
  Fake :: "(state*state) set"
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    23
    "Fake == {(s,s').
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    24
	      EX B X. s' = Says Spy B X # s
5460
0c4e3d024ec9 deleted not_Says_to_self
paulson
parents: 5430
diff changeset
    25
		    & X: synth (analz (spies s))}"
5430
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    26
  
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    27
  (*The numeric suffixes on A identify the rule*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    28
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    29
  (*Alice initiates a protocol run, sending a nonce to Bob*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    30
  NS1 :: "(state*state) set"
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    31
    "NS1 == {(s1,s').
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    32
	     EX A1 B NA.
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    33
	         s' = Says A1 B (Crypt (pubK B) {|Nonce NA, Agent A1|}) # s1
5460
0c4e3d024ec9 deleted not_Says_to_self
paulson
parents: 5430
diff changeset
    34
	       & Nonce NA ~: used s1}"
5430
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    35
  
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    36
  (*Bob responds to Alice's message with a further nonce*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    37
  NS2 :: "(state*state) set"
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    38
    "NS2 == {(s2,s').
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    39
	     EX A' A2 B NA NB.
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    40
	         s' = Says B A2 (Crypt (pubK A2) {|Nonce NA, Nonce NB|}) # s2
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    41
               & Says A' B (Crypt (pubK B) {|Nonce NA, Agent A2|}) : set s2
5460
0c4e3d024ec9 deleted not_Says_to_self
paulson
parents: 5430
diff changeset
    42
	       & Nonce NB ~: used s2}"
5430
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    43
 
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    44
  (*Alice proves her existence by sending NB back to Bob.*)
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    45
  NS3 :: "(state*state) set"
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    46
    "NS3 == {(s3,s').
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    47
	     EX A3 B' B NA NB.
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    48
	         s' = Says A3 B (Crypt (pubK B) (Nonce NB)) # s3
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    49
               & Says A3  B (Crypt (pubK B) {|Nonce NA, Agent A3|}) : set s3
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    50
	       & Says B' A3 (Crypt (pubK A3) {|Nonce NA, Nonce NB|}) : set s3}"
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    51
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    52
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    53
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    54
constdefs
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    55
  Nprg :: state program
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    56
    (*Initial trace is empty*)
6012
1894bfc4aee9 Addition of the States component; parts of Comp not working
paulson
parents: 5596
diff changeset
    57
    "Nprg == mk_program(UNIV, {[]}, {Fake, NS1, NS2, NS3})"
5430
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    58
4a179dba527a New UNITY theory, the N-S protocol
paulson
parents:
diff changeset
    59
end