src/HOL/Auth/Event.thy
author paulson
Fri, 21 Nov 2008 13:17:43 +0100
changeset 28869 191cbfac6c9a
parent 27225 b316dde851f5
child 30510 4120fc59dd85
permissions -rw-r--r--
Strange. The proof worked in the 2008 release. In order to make it work now, the last line of the proof must be moved up two places. In other words, the first proof step is now returning its subgoals in a different order from before.
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
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
     6
Datatype of events; function "spies"; freshness
3678
414e04d7c2d6 Spy can see Notes of the compromised agents
paulson
parents: 3519
diff changeset
     7
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
     8
"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
     9
    stores are visible to him
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    10
*)
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    11
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13935
diff changeset
    12
header{*Theory of Events for Security Protocols*}
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13935
diff changeset
    13
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14200
diff changeset
    14
theory Event imports Message begin
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    15
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    16
consts  (*Initial states of agents -- parameter of the construction*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    17
  initState :: "agent => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    18
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    19
datatype
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    20
  event = Says  agent agent msg
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    21
        | Gets  agent       msg
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    22
        | Notes agent       msg
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    23
       
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    24
consts 
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    25
  bad    :: "agent set"				(*compromised agents*)
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    26
  knows  :: "agent => event list => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    27
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    28
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    29
text{*The constant "spies" is retained for compatibility's sake*}
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18570
diff changeset
    30
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18570
diff changeset
    31
abbreviation (input)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
    32
  spies  :: "event list => msg set" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 18570
diff changeset
    33
  "spies == knows Spy"
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    34
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    35
text{*Spy has access to his own key for spoof messages, but Server is secure*}
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    36
specification (bad)
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
    37
  Spy_in_bad     [iff]: "Spy \<in> bad"
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
    38
  Server_not_bad [iff]: "Server \<notin> bad"
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    39
    by (rule exI [of _ "{Spy}"], simp)
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    40
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 3683
diff changeset
    41
primrec
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    42
  knows_Nil:   "knows A [] = initState A"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    43
  knows_Cons:
6399
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    44
    "knows A (ev # evs) =
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    45
       (if A = Spy then 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    46
	(case ev of
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    47
	   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
    48
	 | Gets A' X => knows Spy evs
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    49
	 | Notes A' X  => 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
    50
	     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
    51
	else
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    52
	(case ev of
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    53
	   Says A' B X => 
4a9040b85e2e exchanged the order of Gets and Notes in datatype event
paulson
parents: 6308
diff changeset
    54
	     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
    55
	 | Gets A' 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
	 | Notes 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))"
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    59
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    60
(*
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    61
  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
    62
  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
    63
  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
    64
*)
3678
414e04d7c2d6 Spy can see Notes of the compromised agents
paulson
parents: 3519
diff changeset
    65
3683
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    66
consts
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    67
  (*Set of items that might be visible to somebody:
aafe719dff14 Global change: lost->bad and sees Spy->spies
paulson
parents: 3678
diff changeset
    68
    complement of the set of fresh items*)
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    69
  used :: "event list => msg set"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
    70
5183
89f162de39cf Adapted to new datatype package.
berghofe
parents: 3683
diff changeset
    71
primrec
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    72
  used_Nil:   "used []         = (UN B. parts (initState B))"
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    73
  used_Cons:  "used (ev # evs) =
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    74
		     (case ev of
13935
paulson
parents: 13926
diff changeset
    75
			Says A B X => parts {X} \<union> used evs
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
    76
		      | Gets A X   => used evs
13935
paulson
parents: 13926
diff changeset
    77
		      | Notes A X  => parts {X} \<union> used evs)"
paulson
parents: 13926
diff changeset
    78
    --{*The case for @{term Gets} seems anomalous, but @{term Gets} always
paulson
parents: 13926
diff changeset
    79
        follows @{term Says} in real protocols.  Seems difficult to change.
paulson
parents: 13926
diff changeset
    80
        See @{text Gets_correct} in theory @{text "Guard/Extensions.thy"}. *}
6308
76f3865a2b1d Added Bella's "Gets" model for Otway_Rees. Also affects some other theories.
paulson
parents: 5183
diff changeset
    81
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    82
lemma Notes_imp_used [rule_format]: "Notes A X \<in> set evs --> X \<in> used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    83
apply (induct_tac evs)
11463
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
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    87
lemma Says_imp_used [rule_format]: "Says A B X \<in> set evs --> X \<in> used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    88
apply (induct_tac evs)
11463
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
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    92
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    93
subsection{*Function @{term knows}*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    94
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13935
diff changeset
    95
(*Simplifying   
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13935
diff changeset
    96
 parts(insert X (knows Spy evs)) = parts{X} \<union> parts(knows Spy evs).
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13935
diff changeset
    97
  This version won't loop with the simplifier.*)
13935
paulson
parents: 13926
diff changeset
    98
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs", standard]
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    99
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   100
lemma knows_Spy_Says [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   101
     "knows Spy (Says A B X # evs) = insert X (knows Spy evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   102
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   103
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
   104
text{*Letting the Spy see "bad" agents' notes avoids redundant case-splits
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14126
diff changeset
   105
      on whether @{term "A=Spy"} and whether @{term "A\<in>bad"}*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   106
lemma knows_Spy_Notes [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   107
     "knows Spy (Notes A X # evs) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   108
          (if A:bad then insert X (knows Spy evs) else knows Spy evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   109
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   110
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   111
lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   112
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   113
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   114
lemma knows_Spy_subset_knows_Spy_Says:
13935
paulson
parents: 13926
diff changeset
   115
     "knows Spy evs \<subseteq> knows Spy (Says A B X # evs)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   116
by (simp add: subset_insertI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   117
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   118
lemma knows_Spy_subset_knows_Spy_Notes:
13935
paulson
parents: 13926
diff changeset
   119
     "knows Spy evs \<subseteq> knows Spy (Notes A X # evs)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   120
by force
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   121
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   122
lemma knows_Spy_subset_knows_Spy_Gets:
13935
paulson
parents: 13926
diff changeset
   123
     "knows Spy evs \<subseteq> knows Spy (Gets A X # evs)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   124
by (simp add: subset_insertI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   125
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   126
text{*Spy sees what is sent on the traffic*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   127
lemma Says_imp_knows_Spy [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   128
     "Says A B X \<in> set evs --> X \<in> knows Spy evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   129
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   130
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   131
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   132
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   133
lemma Notes_imp_knows_Spy [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   134
     "Notes A X \<in> set evs --> A: bad --> X \<in> knows Spy evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   135
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   136
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   137
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   138
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   139
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   140
text{*Elimination rules: derive contradictions from old Says events containing
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   141
  items known to be fresh*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   142
lemmas knows_Spy_partsEs =
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   143
     Says_imp_knows_Spy [THEN parts.Inj, THEN revcut_rl, standard] 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   144
     parts.Body [THEN revcut_rl, standard]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   145
18570
ffce25f9aa7f a few more named lemmas
paulson
parents: 17990
diff changeset
   146
lemmas Says_imp_analz_Spy = Says_imp_knows_Spy [THEN analz.Inj]
ffce25f9aa7f a few more named lemmas
paulson
parents: 17990
diff changeset
   147
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   148
text{*Compatibility for the old "spies" function*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   149
lemmas spies_partsEs = knows_Spy_partsEs
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   150
lemmas Says_imp_spies = Says_imp_knows_Spy
13935
paulson
parents: 13926
diff changeset
   151
lemmas parts_insert_spies = parts_insert_knows_A [of _ Spy]
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   152
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   153
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   154
subsection{*Knowledge of Agents*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   155
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   156
lemma knows_Says: "knows A (Says A B X # evs) = insert X (knows A evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   157
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   158
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   159
lemma knows_Notes: "knows A (Notes A X # evs) = insert X (knows A evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   160
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   161
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   162
lemma knows_Gets:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   163
     "A \<noteq> Spy --> knows A (Gets A X # evs) = insert X (knows A evs)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   164
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   165
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   166
13935
paulson
parents: 13926
diff changeset
   167
lemma knows_subset_knows_Says: "knows A evs \<subseteq> knows A (Says A' B X # evs)"
paulson
parents: 13926
diff changeset
   168
by (simp add: subset_insertI)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   169
13935
paulson
parents: 13926
diff changeset
   170
lemma knows_subset_knows_Notes: "knows A evs \<subseteq> knows A (Notes A' X # evs)"
paulson
parents: 13926
diff changeset
   171
by (simp add: subset_insertI)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   172
13935
paulson
parents: 13926
diff changeset
   173
lemma knows_subset_knows_Gets: "knows A evs \<subseteq> knows A (Gets A' X # evs)"
paulson
parents: 13926
diff changeset
   174
by (simp add: subset_insertI)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   175
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   176
text{*Agents know what they say*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   177
lemma Says_imp_knows [rule_format]: "Says A B X \<in> set evs --> X \<in> knows A evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   178
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   179
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   180
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   181
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   182
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   183
text{*Agents know what they note*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   184
lemma Notes_imp_knows [rule_format]: "Notes A X \<in> set evs --> X \<in> knows A evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   185
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   186
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   187
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   188
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   189
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   190
text{*Agents know what they receive*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   191
lemma Gets_imp_knows_agents [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   192
     "A \<noteq> Spy --> Gets A X \<in> set evs --> X \<in> knows A evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   193
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   194
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   195
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   196
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   197
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   198
text{*What agents DIFFERENT FROM Spy know 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   199
  was either said, or noted, or got, or known initially*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   200
lemma knows_imp_Says_Gets_Notes_initState [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   201
     "[| X \<in> knows A evs; A \<noteq> Spy |] ==> EX B.  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   202
  Says A B X \<in> set evs | Gets A X \<in> set evs | Notes A X \<in> set evs | X \<in> initState A"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   203
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   204
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   205
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   206
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   207
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   208
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   209
text{*What the Spy knows -- for the time being --
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   210
  was either said or noted, or known initially*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   211
lemma knows_Spy_imp_Says_Notes_initState [rule_format]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   212
     "[| X \<in> knows Spy evs |] ==> EX A B.  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   213
  Says A B X \<in> set evs | Notes A X \<in> set evs | X \<in> initState Spy"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   214
apply (erule rev_mp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   215
apply (induct_tac "evs")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   216
apply (simp_all (no_asm_simp) split add: event.split)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   217
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   218
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   219
13935
paulson
parents: 13926
diff changeset
   220
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs"
paulson
parents: 13926
diff changeset
   221
apply (induct_tac "evs", force)  
paulson
parents: 13926
diff changeset
   222
apply (simp add: parts_insert_knows_A knows_Cons add: event.split, blast) 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   223
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   224
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   225
lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   226
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   227
lemma initState_into_used: "X \<in> parts (initState B) ==> X \<in> used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   228
apply (induct_tac "evs")
13935
paulson
parents: 13926
diff changeset
   229
apply (simp_all add: parts_insert_knows_A split add: event.split, blast)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   230
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   231
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   232
lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   233
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   234
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   235
lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   236
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   237
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   238
lemma used_Gets [simp]: "used (Gets A X # evs) = used evs"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   239
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   240
13935
paulson
parents: 13926
diff changeset
   241
lemma used_nil_subset: "used [] \<subseteq> used evs"
paulson
parents: 13926
diff changeset
   242
apply simp
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   243
apply (blast intro: initState_into_used)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   244
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   245
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   246
text{*NOTE REMOVAL--laws above are cleaner, as they don't involve "case"*}
13935
paulson
parents: 13926
diff changeset
   247
declare knows_Cons [simp del]
paulson
parents: 13926
diff changeset
   248
        used_Nil [simp del] used_Cons [simp del]
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   249
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   250
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   251
text{*For proving theorems of the form @{term "X \<notin> analz (knows Spy evs) --> P"}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   252
  New events added by induction to "evs" are discarded.  Provided 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   253
  this information isn't needed, the proof will be much shorter, since
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   254
  it will omit complicated reasoning about @{term analz}.*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   255
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   256
lemmas analz_mono_contra =
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   257
       knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   258
       knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   259
       knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   260
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
   261
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   262
lemma knows_subset_knows_Cons: "knows A evs \<subseteq> knows A (e # evs)"
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   263
by (induct e, auto simp: knows_Cons)
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   264
13935
paulson
parents: 13926
diff changeset
   265
lemma initState_subset_knows: "initState A \<subseteq> knows A evs"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   266
apply (induct_tac evs, simp) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   267
apply (blast intro: knows_subset_knows_Cons [THEN subsetD])
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   268
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   269
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   270
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   271
text{*For proving @{text new_keys_not_used}*}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   272
lemma keysFor_parts_insert:
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   273
     "[| K \<in> keysFor (parts (insert X G));  X \<in> synth (analz H) |] 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   274
      ==> K \<in> keysFor (parts (G \<union> H)) | Key (invKey K) \<in> parts H"; 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   275
by (force 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   276
    dest!: parts_insert_subset_Un [THEN keysFor_mono, THEN [2] rev_subsetD]
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   277
           analz_subset_parts [THEN keysFor_mono, THEN [2] rev_subsetD]
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   278
    intro: analz_subset_parts [THEN subsetD] parts_mono [THEN [2] rev_subsetD])
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   279
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   280
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   281
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)", standard]
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   282
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   283
ML
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   284
{*
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   285
val analz_mono_contra_tac = 
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   286
  rtac @{thm analz_impI} THEN' 
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   287
  REPEAT1 o (dresolve_tac @{thms analz_mono_contra})
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   288
  THEN' mp_tac
24122
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   289
*}
fc7f857d33c8 tuned ML bindings (for multithreading);
wenzelm
parents: 21588
diff changeset
   290
11104
f2024fed9f0c partial conversion to Isar script style
paulson
parents: 6399
diff changeset
   291
method_setup analz_mono_contra = {*
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21404
diff changeset
   292
    Method.no_args (Method.SIMPLE_METHOD (REPEAT_FIRST analz_mono_contra_tac)) *}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   293
    "for proving theorems of the form X \<notin> analz (knows Spy evs) --> P"
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   294
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   295
subsubsection{*Useful for case analysis on whether a hash is a spoof or not*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   296
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   297
lemmas syan_impI = impI [where P = "Y \<notin> synth (analz (knows Spy evs))", standard]
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   298
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   299
ML
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   300
{*
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   301
val synth_analz_mono_contra_tac = 
27225
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   302
  rtac @{thm syan_impI} THEN'
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   303
  REPEAT1 o 
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   304
    (dresolve_tac 
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   305
     [@{thm knows_Spy_subset_knows_Spy_Says} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   306
      @{thm knows_Spy_subset_knows_Spy_Notes} RS @{thm synth_analz_mono} RS @{thm contra_subsetD},
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   307
      @{thm knows_Spy_subset_knows_Spy_Gets} RS @{thm synth_analz_mono} RS @{thm contra_subsetD}])
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   308
  THEN'
b316dde851f5 eliminated OldGoals.inst;
wenzelm
parents: 27154
diff changeset
   309
  mp_tac
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   310
*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   311
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   312
method_setup synth_analz_mono_contra = {*
21588
cd0dc678a205 simplified method setup;
wenzelm
parents: 21404
diff changeset
   313
    Method.no_args (Method.SIMPLE_METHOD (REPEAT_FIRST synth_analz_mono_contra_tac)) *}
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11463
diff changeset
   314
    "for proving theorems of the form X \<notin> synth (analz (knows Spy evs)) --> P"
3512
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
   315
9dcb4daa15e8 Moving common declarations and proofs from theories "Shared"
paulson
parents:
diff changeset
   316
end