src/HOL/Auth/Event.thy
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 11463 96b5b27da55c
child 13922 75ae4244a596
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Event
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     2
    ID:         $Id$
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     5
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     6
Theory of events for security protocols
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
     7
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
     8
Datatype of events; function "spies"; freshness
3678
414e04d7c2d6 Spy can see Notes of the compromised agents
paulson
parents: 3519
diff changeset
     9
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    10
"bad" agents have been broken by the Spy; their private keys and internal
3678
414e04d7c2d6 Spy can see Notes of the compromised agents
paulson
parents: 3519
diff changeset
    11
    stores are visible to him
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    12
*)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    13
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    14
theory Event = Message
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    15
files ("Event_lemmas.ML"):
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    16
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    17
consts  (*Initial states of agents -- parameter of the construction*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    18
  initState :: "agent => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    19
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    20
datatype
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    21
  event = Says  agent agent msg
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    22
        | Gets  agent       msg
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    23
        | Notes agent       msg
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    24
       
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    25
consts 
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    26
  bad    :: "agent set"				(*compromised agents*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    27
  knows  :: "agent => event list => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    28
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    29
11310
51e70b7bc315 spelling check
paulson
parents: 11192
diff changeset
    30
(*"spies" is retained for compatibility's sake*)
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    31
syntax
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    32
  spies  :: "event list => msg set"
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    33
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    34
translations
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    35
  "spies"   => "knows Spy"
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    36
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    37
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    38
axioms
3519
ab0a9fbed4c0 Changing "lost" from a parameter of protocol definitions to a constant.
paulson
parents: 3512
diff changeset
    39
  (*Spy has access to his own key for spoof messages, but Server is secure*)
11310
51e70b7bc315 spelling check
paulson
parents: 11192
diff changeset
    40
  Spy_in_bad     [iff] :    "Spy \\<in> bad"
51e70b7bc315 spelling check
paulson
parents: 11192
diff changeset
    41
  Server_not_bad [iff] : "Server \\<notin> bad"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    42
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 3683
diff changeset
    43
primrec
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    44
  knows_Nil:   "knows A [] = initState A"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    45
  knows_Cons:
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    46
    "knows A (ev # evs) =
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    47
       (if A = Spy then 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    48
	(case ev of
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    49
	   Says A' B X => insert X (knows Spy evs)
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    50
	 | Gets A' X => knows Spy evs
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    51
	 | Notes A' X  => 
11310
51e70b7bc315 spelling check
paulson
parents: 11192
diff changeset
    52
	     if A' \\<in> bad then insert X (knows Spy evs) else knows Spy evs)
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    53
	else
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    54
	(case ev of
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    55
	   Says A' B X => 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    56
	     if A'=A then insert X (knows A evs) else knows A evs
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    57
	 | Gets A' X    => 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    58
	     if A'=A then insert X (knows A evs) else knows A evs
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    59
	 | Notes A' X    => 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    60
	     if A'=A then insert X (knows A evs) else knows A evs))"
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    61
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    62
(*
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    63
  Case A=Spy on the Gets event
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    64
  enforces the fact that if a message is received then it must have been sent,
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    65
  therefore the oops case must use Notes
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    66
*)
3678
414e04d7c2d6 Spy can see Notes of the compromised agents
paulson
parents: 3519
diff changeset
    67
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    68
consts
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    69
  (*Set of items that might be visible to somebody:
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    70
    complement of the set of fresh items*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    71
  used :: "event list => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    72
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 3683
diff changeset
    73
primrec
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    74
  used_Nil:   "used []         = (UN B. parts (initState B))"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    75
  used_Cons:  "used (ev # evs) =
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    76
		     (case ev of
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    77
			Says A B X => parts {X} Un (used evs)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    78
		      | Gets A X   => used evs
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    79
		      | Notes A X  => parts {X} Un (used evs))"
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    80
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    81
11463
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    82
lemma Notes_imp_used [rule_format]: "Notes A X : set evs --> X : used evs"
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    83
apply (induct_tac evs);
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    84
apply (auto split: event.split) 
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    85
done
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    86
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    87
lemma Says_imp_used [rule_format]: "Says A B X : set evs --> X : used evs"
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    88
apply (induct_tac evs);
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    89
apply (auto split: event.split) 
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    90
done
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    91
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    92
lemma MPair_used [rule_format]:
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    93
     "MPair X Y : used evs --> X : used evs & Y : used evs"
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    94
apply (induct_tac evs);
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    95
apply (auto split: event.split) 
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    96
done
96b5b27da55c three new theorems
paulson
parents: 11310
diff changeset
    97
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    98
use "Event_lemmas.ML"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    99
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
   100
method_setup analz_mono_contra = {*
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
   101
    Method.no_args
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
   102
      (Method.METHOD (fn facts => REPEAT_FIRST analz_mono_contra_tac)) *}
11310
51e70b7bc315 spelling check
paulson
parents: 11192
diff changeset
   103
    "for proving theorems of the form X \\<notin> analz (knows Spy evs) --> P"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
   104
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
   105
end