src/HOL/Auth/Message.thy
author paulson
Tue, 28 Mar 2006 16:48:18 +0200
changeset 19334 96ca738055a6
parent 18492 b0fe60800623
child 20648 742c30fc3fcb
permissions -rw-r--r--
Simplified version of Jia's filter. Now all constants are pooled, rather than relevance being compared against separate clauses. Rejects are no longer noted, and units cannot be added at the end.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Message
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     2
    ID:         $Id$
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     5
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     6
Datatypes of agents and messages;
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1839
diff changeset
     7
Inductive relations "parts", "analz" and "synth"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     8
*)
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
     9
13956
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13926
diff changeset
    10
header{*Theory of Agents and Messages for Security Protocols*}
8fe7e12290e1 improved presentation of HOL/Auth theories
paulson
parents: 13926
diff changeset
    11
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15032
diff changeset
    12
theory Message imports Main begin
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    13
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    14
(*Needed occasionally with spy_analz_tac, e.g. in analz_insert_Key_newK*)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    15
lemma [simp] : "A \<union> (B \<union> A) = B \<union> A"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    16
by blast
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    17
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    18
types 
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    19
  key = nat
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    20
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    21
consts
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    22
  all_symmetric :: bool        --{*true if all keys are symmetric*}
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    23
  invKey        :: "key=>key"  --{*inverse of a symmetric key*}
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    24
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    25
specification (invKey)
14181
942db403d4bb new, separate specifications
paulson
parents: 14145
diff changeset
    26
  invKey [simp]: "invKey (invKey K) = K"
942db403d4bb new, separate specifications
paulson
parents: 14145
diff changeset
    27
  invKey_symmetric: "all_symmetric --> invKey = id"
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    28
    by (rule exI [of _ id], auto)
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    29
14126
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    30
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    31
text{*The inverse of a symmetric key is itself; that of a public key
28824746d046 Tidying and replacement of some axioms by specifications
paulson
parents: 13956
diff changeset
    32
      is the private key and vice versa*}
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    33
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    34
constdefs
11230
756c5034f08b misc tidying; changing the predicate isSymKey to the set symKeys
paulson
parents: 11192
diff changeset
    35
  symKeys :: "key set"
756c5034f08b misc tidying; changing the predicate isSymKey to the set symKeys
paulson
parents: 11192
diff changeset
    36
  "symKeys == {K. invKey K = K}"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    37
16818
paulson
parents: 16796
diff changeset
    38
datatype  --{*We allow any number of friendly agents*}
2032
1bbf1bdcaf56 Introduction of "lost" argument
paulson
parents: 2010
diff changeset
    39
  agent = Server | Friend nat | Spy
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    40
3668
a39baf59ea47 Split base cases from "msg" to "atomic" in order
paulson
parents: 2516
diff changeset
    41
datatype
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    42
     msg = Agent  agent	    --{*Agent names*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    43
         | Number nat       --{*Ordinary integers, timestamps, ...*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    44
         | Nonce  nat       --{*Unguessable nonces*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    45
         | Key    key       --{*Crypto keys*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    46
	 | Hash   msg       --{*Hashing*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    47
	 | MPair  msg msg   --{*Compound messages*}
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
    48
	 | Crypt  key msg   --{*Encryption, public- or shared-key*}
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    49
5234
701fa0ed77b7 Better comments
paulson
parents: 5183
diff changeset
    50
16818
paulson
parents: 16796
diff changeset
    51
text{*Concrete syntax: messages appear as {|A,B,NA|}, etc...*}
5234
701fa0ed77b7 Better comments
paulson
parents: 5183
diff changeset
    52
syntax
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2484
diff changeset
    53
  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2{|_,/ _|})")
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    54
9686
87b460d72e80 xsymbols for {| and |}
paulson
parents: 7057
diff changeset
    55
syntax (xsymbols)
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    56
  "@MTuple"      :: "['a, args] => 'a * 'b"       ("(2\<lbrace>_,/ _\<rbrace>)")
9686
87b460d72e80 xsymbols for {| and |}
paulson
parents: 7057
diff changeset
    57
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    58
translations
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    59
  "{|x, y, z|}"   == "{|x, {|y, z|}|}"
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    60
  "{|x, y|}"      == "MPair x y"
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    61
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    62
2484
596a5b5a68ff Incorporation of HPair into Message
paulson
parents: 2373
diff changeset
    63
constdefs
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    64
  HPair :: "[msg,msg] => msg"                       ("(4Hash[_] /_)" [0, 1000])
16818
paulson
parents: 16796
diff changeset
    65
    --{*Message Y paired with a MAC computed with the help of X*}
2516
4d68fbe6378b Now with Andy Gordon's treatment of freshness to replace newN/K
paulson
parents: 2484
diff changeset
    66
    "Hash[X] Y == {| Hash{|X,Y|}, Y|}"
2484
596a5b5a68ff Incorporation of HPair into Message
paulson
parents: 2373
diff changeset
    67
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    68
  keysFor :: "msg set => key set"
16818
paulson
parents: 16796
diff changeset
    69
    --{*Keys useful to decrypt elements of a message set*}
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
    70
  "keysFor H == invKey ` {K. \<exists>X. Crypt K X \<in> H}"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    71
16818
paulson
parents: 16796
diff changeset
    72
paulson
parents: 16796
diff changeset
    73
subsubsection{*Inductive Definition of All Parts" of a Message*}
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    74
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    75
consts  parts   :: "msg set => msg set"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    76
inductive "parts H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    77
  intros 
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
    78
    Inj [intro]:               "X \<in> H ==> X \<in> parts H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
    79
    Fst:         "{|X,Y|}   \<in> parts H ==> X \<in> parts H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
    80
    Snd:         "{|X,Y|}   \<in> parts H ==> Y \<in> parts H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
    81
    Body:        "Crypt K X \<in> parts H ==> X \<in> parts H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    82
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    83
16818
paulson
parents: 16796
diff changeset
    84
text{*Monotonicity*}
paulson
parents: 16796
diff changeset
    85
lemma parts_mono: "G \<subseteq> H ==> parts(G) \<subseteq> parts(H)"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    86
apply auto
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    87
apply (erule parts.induct) 
16818
paulson
parents: 16796
diff changeset
    88
apply (blast dest: parts.Fst parts.Snd parts.Body)+
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
    89
done
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    90
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
    91
16818
paulson
parents: 16796
diff changeset
    92
text{*Equations hold because constructors are injective.*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    93
lemma Friend_image_eq [simp]: "(Friend x \<in> Friend`A) = (x:A)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    94
by auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    95
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    96
lemma Key_image_eq [simp]: "(Key x \<in> Key`A) = (x\<in>A)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    97
by auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    98
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
    99
lemma Nonce_Key_image_eq [simp]: "(Nonce x \<notin> Key`A)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   100
by auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   101
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   102
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   103
subsubsection{*Inverse of keys *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   104
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   105
lemma invKey_eq [simp]: "(invKey K = invKey K') = (K=K')"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   106
apply safe
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   107
apply (drule_tac f = invKey in arg_cong, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   108
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   109
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
subsection{*keysFor operator*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   112
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   113
lemma keysFor_empty [simp]: "keysFor {} = {}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   114
by (unfold keysFor_def, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   115
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   116
lemma keysFor_Un [simp]: "keysFor (H \<union> H') = keysFor H \<union> keysFor H'"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   117
by (unfold keysFor_def, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   118
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   119
lemma keysFor_UN [simp]: "keysFor (\<Union>i\<in>A. H i) = (\<Union>i\<in>A. keysFor (H i))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   120
by (unfold keysFor_def, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   121
16818
paulson
parents: 16796
diff changeset
   122
text{*Monotonicity*}
paulson
parents: 16796
diff changeset
   123
lemma keysFor_mono: "G \<subseteq> H ==> keysFor(G) \<subseteq> keysFor(H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   124
by (unfold keysFor_def, blast)
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
lemma keysFor_insert_Agent [simp]: "keysFor (insert (Agent A) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   127
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   128
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   129
lemma keysFor_insert_Nonce [simp]: "keysFor (insert (Nonce N) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   130
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   131
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   132
lemma keysFor_insert_Number [simp]: "keysFor (insert (Number N) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   133
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   134
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   135
lemma keysFor_insert_Key [simp]: "keysFor (insert (Key K) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   136
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   137
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   138
lemma keysFor_insert_Hash [simp]: "keysFor (insert (Hash X) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   139
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   140
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   141
lemma keysFor_insert_MPair [simp]: "keysFor (insert {|X,Y|} H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   142
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   143
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   144
lemma keysFor_insert_Crypt [simp]: 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   145
    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)"
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   146
by (unfold keysFor_def, auto)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   147
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   148
lemma keysFor_image_Key [simp]: "keysFor (Key`E) = {}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   149
by (unfold keysFor_def, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   150
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   151
lemma Crypt_imp_invKey_keysFor: "Crypt K X \<in> H ==> invKey K \<in> keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   152
by (unfold keysFor_def, blast)
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
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   155
subsection{*Inductive relation "parts"*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   156
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   157
lemma MPair_parts:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   158
     "[| {|X,Y|} \<in> parts H;        
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   159
         [| X \<in> parts H; Y \<in> parts H |] ==> P |] ==> P"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   160
by (blast dest: parts.Fst parts.Snd) 
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
declare MPair_parts [elim!]  parts.Body [dest!]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   163
text{*NB These two rules are UNSAFE in the formal sense, as they discard the
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   164
     compound message.  They work well on THIS FILE.  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   165
  @{text MPair_parts} is left as SAFE because it speeds up proofs.
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   166
  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   167
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   168
lemma parts_increasing: "H \<subseteq> parts(H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   169
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   170
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   171
lemmas parts_insertI = subset_insertI [THEN parts_mono, THEN subsetD, standard]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   172
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   173
lemma parts_empty [simp]: "parts{} = {}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   174
apply safe
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   175
apply (erule parts.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   176
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   177
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   178
lemma parts_emptyE [elim!]: "X\<in> parts{} ==> P"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   179
by simp
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   180
16818
paulson
parents: 16796
diff changeset
   181
text{*WARNING: loops if H = {Y}, therefore must not be repeated!*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   182
lemma parts_singleton: "X\<in> parts H ==> \<exists>Y\<in>H. X\<in> parts {Y}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   183
by (erule parts.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   184
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   185
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   186
subsubsection{*Unions *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   187
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   188
lemma parts_Un_subset1: "parts(G) \<union> parts(H) \<subseteq> parts(G \<union> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   189
by (intro Un_least parts_mono Un_upper1 Un_upper2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   190
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   191
lemma parts_Un_subset2: "parts(G \<union> H) \<subseteq> parts(G) \<union> parts(H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   192
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   193
apply (erule parts.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   194
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   195
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   196
lemma parts_Un [simp]: "parts(G \<union> H) = parts(G) \<union> parts(H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   197
by (intro equalityI parts_Un_subset1 parts_Un_subset2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   198
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   199
lemma parts_insert: "parts (insert X H) = parts {X} \<union> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   200
apply (subst insert_is_Un [of _ H])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   201
apply (simp only: parts_Un)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   202
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   203
16818
paulson
parents: 16796
diff changeset
   204
text{*TWO inserts to avoid looping.  This rewrite is better than nothing.
paulson
parents: 16796
diff changeset
   205
  Not suitable for Addsimps: its behaviour can be strange.*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   206
lemma parts_insert2:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   207
     "parts (insert X (insert Y H)) = parts {X} \<union> parts {Y} \<union> parts H"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   208
apply (simp add: Un_assoc)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   209
apply (simp add: parts_insert [symmetric])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   210
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   211
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   212
lemma parts_UN_subset1: "(\<Union>x\<in>A. parts(H x)) \<subseteq> parts(\<Union>x\<in>A. H x)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   213
by (intro UN_least parts_mono UN_upper)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   214
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   215
lemma parts_UN_subset2: "parts(\<Union>x\<in>A. H x) \<subseteq> (\<Union>x\<in>A. parts(H x))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   216
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   217
apply (erule parts.induct, 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
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   220
lemma parts_UN [simp]: "parts(\<Union>x\<in>A. H x) = (\<Union>x\<in>A. parts(H x))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   221
by (intro equalityI parts_UN_subset1 parts_UN_subset2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   222
16818
paulson
parents: 16796
diff changeset
   223
text{*Added to simplify arguments to parts, analz and synth.
paulson
parents: 16796
diff changeset
   224
  NOTE: the UN versions are no longer used!*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   225
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
text{*This allows @{text blast} to simplify occurrences of 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   228
  @{term "parts(G\<union>H)"} in the assumption.*}
17729
d74d0b5052a0 theorems need names
paulson
parents: 17689
diff changeset
   229
lemmas in_parts_UnE = parts_Un [THEN equalityD1, THEN subsetD, THEN UnE] 
d74d0b5052a0 theorems need names
paulson
parents: 17689
diff changeset
   230
declare in_parts_UnE [elim!]
13926
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
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   233
lemma parts_insert_subset: "insert X (parts H) \<subseteq> parts(insert X H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   234
by (blast intro: parts_mono [THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   235
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   236
subsubsection{*Idempotence and transitivity *}
13926
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 parts_partsD [dest!]: "X\<in> parts (parts H) ==> X\<in> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   239
by (erule parts.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   240
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   241
lemma parts_idem [simp]: "parts (parts H) = parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   242
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   243
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   244
lemma parts_subset_iff [simp]: "(parts G \<subseteq> parts H) = (G \<subseteq> parts H)"
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   245
apply (rule iffI)
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   246
apply (iprover intro: subset_trans parts_increasing)  
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   247
apply (frule parts_mono, simp) 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   248
done
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   249
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   250
lemma parts_trans: "[| X\<in> parts G;  G \<subseteq> parts H |] ==> X\<in> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   251
by (drule parts_mono, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   252
16818
paulson
parents: 16796
diff changeset
   253
text{*Cut*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   254
lemma parts_cut:
18492
b0fe60800623 shorter proof
paulson
parents: 17729
diff changeset
   255
     "[| Y\<in> parts (insert X G);  X\<in> parts H |] ==> Y\<in> parts (G \<union> H)" 
b0fe60800623 shorter proof
paulson
parents: 17729
diff changeset
   256
by (blast intro: parts_trans) 
b0fe60800623 shorter proof
paulson
parents: 17729
diff changeset
   257
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   258
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   259
lemma parts_cut_eq [simp]: "X\<in> parts H ==> parts (insert X H) = parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   260
by (force dest!: parts_cut intro: parts_insertI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   261
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   262
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   263
subsubsection{*Rewrite rules for pulling out atomic messages *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   264
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   265
lemmas parts_insert_eq_I = equalityI [OF subsetI parts_insert_subset]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   266
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   267
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   268
lemma parts_insert_Agent [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   269
     "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   270
apply (rule parts_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   271
apply (erule parts.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   272
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   273
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   274
lemma parts_insert_Nonce [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   275
     "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   276
apply (rule parts_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   277
apply (erule parts.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   278
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   279
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   280
lemma parts_insert_Number [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   281
     "parts (insert (Number N) H) = insert (Number N) (parts H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   282
apply (rule parts_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   283
apply (erule parts.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   284
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   285
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   286
lemma parts_insert_Key [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   287
     "parts (insert (Key K) H) = insert (Key K) (parts H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   288
apply (rule parts_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   289
apply (erule parts.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   290
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   291
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   292
lemma parts_insert_Hash [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   293
     "parts (insert (Hash X) H) = insert (Hash X) (parts H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   294
apply (rule parts_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   295
apply (erule parts.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   296
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   297
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   298
lemma parts_insert_Crypt [simp]:
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   299
     "parts (insert (Crypt K X) H) = insert (Crypt K X) (parts (insert X H))"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   300
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   301
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   302
apply (erule parts.induct, auto)
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   303
apply (blast intro: parts.Body)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   304
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   305
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   306
lemma parts_insert_MPair [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   307
     "parts (insert {|X,Y|} H) =  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   308
          insert {|X,Y|} (parts (insert X (insert Y H)))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   309
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   310
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   311
apply (erule parts.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   312
apply (blast intro: parts.Fst parts.Snd)+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   313
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   314
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   315
lemma parts_image_Key [simp]: "parts (Key`N) = Key`N"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   316
apply auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   317
apply (erule parts.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   318
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   319
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   320
16818
paulson
parents: 16796
diff changeset
   321
text{*In any message, there is an upper bound N on its greatest nonce.*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   322
lemma msg_Nonce_supply: "\<exists>N. \<forall>n. N\<le>n --> Nonce n \<notin> parts {msg}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   323
apply (induct_tac "msg")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   324
apply (simp_all (no_asm_simp) add: exI parts_insert2)
16818
paulson
parents: 16796
diff changeset
   325
 txt{*MPair case: blast works out the necessary sum itself!*}
paulson
parents: 16796
diff changeset
   326
 prefer 2 apply (blast elim!: add_leE)
paulson
parents: 16796
diff changeset
   327
txt{*Nonce case*}
paulson
parents: 16796
diff changeset
   328
apply (rule_tac x = "N + Suc nat" in exI, auto) 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   329
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   330
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   331
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   332
subsection{*Inductive relation "analz"*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   333
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   334
text{*Inductive definition of "analz" -- what can be broken down from a set of
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   335
    messages, including keys.  A form of downward closure.  Pairs can
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   336
    be taken apart; messages decrypted with known keys.  *}
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   337
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   338
consts  analz   :: "msg set => msg set"
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1839
diff changeset
   339
inductive "analz H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   340
  intros 
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   341
    Inj [intro,simp] :    "X \<in> H ==> X \<in> analz H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   342
    Fst:     "{|X,Y|} \<in> analz H ==> X \<in> analz H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   343
    Snd:     "{|X,Y|} \<in> analz H ==> Y \<in> analz H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   344
    Decrypt [dest]: 
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   345
             "[|Crypt K X \<in> analz H; Key(invKey K): analz H|] ==> X \<in> analz H"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   346
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   347
16818
paulson
parents: 16796
diff changeset
   348
text{*Monotonicity; Lemma 1 of Lowe's paper*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   349
lemma analz_mono: "G\<subseteq>H ==> analz(G) \<subseteq> analz(H)"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   350
apply auto
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   351
apply (erule analz.induct) 
16818
paulson
parents: 16796
diff changeset
   352
apply (auto dest: analz.Fst analz.Snd) 
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   353
done
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   354
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   355
text{*Making it safe speeds up proofs*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   356
lemma MPair_analz [elim!]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   357
     "[| {|X,Y|} \<in> analz H;        
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   358
             [| X \<in> analz H; Y \<in> analz H |] ==> P   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   359
          |] ==> P"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   360
by (blast dest: analz.Fst analz.Snd)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   361
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   362
lemma analz_increasing: "H \<subseteq> analz(H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   363
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   364
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   365
lemma analz_subset_parts: "analz H \<subseteq> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   366
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   367
apply (erule analz.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   368
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   369
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   370
lemmas analz_into_parts = analz_subset_parts [THEN subsetD, standard]
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   371
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   372
lemmas not_parts_not_analz = analz_subset_parts [THEN contra_subsetD, standard]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   373
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   374
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   375
lemma parts_analz [simp]: "parts (analz H) = parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   376
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   377
apply (rule analz_subset_parts [THEN parts_mono, THEN subset_trans], simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   378
apply (blast intro: analz_increasing [THEN parts_mono, THEN subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   379
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   380
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   381
lemma analz_parts [simp]: "analz (parts H) = parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   382
apply auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   383
apply (erule analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   384
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   385
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   386
lemmas analz_insertI = subset_insertI [THEN analz_mono, THEN [2] rev_subsetD, standard]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   387
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   388
subsubsection{*General equational properties *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   389
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   390
lemma analz_empty [simp]: "analz{} = {}"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   391
apply safe
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   392
apply (erule analz.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   393
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   394
16818
paulson
parents: 16796
diff changeset
   395
text{*Converse fails: we can analz more from the union than from the 
paulson
parents: 16796
diff changeset
   396
  separate parts, as a key in one might decrypt a message in the other*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   397
lemma analz_Un: "analz(G) \<union> analz(H) \<subseteq> analz(G \<union> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   398
by (intro Un_least analz_mono Un_upper1 Un_upper2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   399
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   400
lemma analz_insert: "insert X (analz H) \<subseteq> analz(insert X H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   401
by (blast intro: analz_mono [THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   402
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   403
subsubsection{*Rewrite rules for pulling out atomic messages *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   404
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   405
lemmas analz_insert_eq_I = equalityI [OF subsetI analz_insert]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   406
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   407
lemma analz_insert_Agent [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   408
     "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   409
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   410
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   411
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   412
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   413
lemma analz_insert_Nonce [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   414
     "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   415
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   416
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   417
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   418
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   419
lemma analz_insert_Number [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   420
     "analz (insert (Number N) H) = insert (Number N) (analz H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   421
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   422
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   423
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   424
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   425
lemma analz_insert_Hash [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   426
     "analz (insert (Hash X) H) = insert (Hash X) (analz H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   427
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   428
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   429
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   430
16818
paulson
parents: 16796
diff changeset
   431
text{*Can only pull out Keys if they are not needed to decrypt the rest*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   432
lemma analz_insert_Key [simp]: 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   433
    "K \<notin> keysFor (analz H) ==>   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   434
          analz (insert (Key K) H) = insert (Key K) (analz H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   435
apply (unfold keysFor_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   436
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   437
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   438
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   439
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   440
lemma analz_insert_MPair [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   441
     "analz (insert {|X,Y|} H) =  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   442
          insert {|X,Y|} (analz (insert X (insert Y H)))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   443
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   444
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   445
apply (erule analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   446
apply (erule analz.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   447
apply (blast intro: analz.Fst analz.Snd)+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   448
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   449
16818
paulson
parents: 16796
diff changeset
   450
text{*Can pull out enCrypted message if the Key is not known*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   451
lemma analz_insert_Crypt:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   452
     "Key (invKey K) \<notin> analz H 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   453
      ==> analz (insert (Crypt K X) H) = insert (Crypt K X) (analz H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   454
apply (rule analz_insert_eq_I) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   455
apply (erule analz.induct, auto) 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   456
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   457
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   458
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   459
lemma lemma1: "Key (invKey K) \<in> analz H ==>   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   460
               analz (insert (Crypt K X) H) \<subseteq>  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   461
               insert (Crypt K X) (analz (insert X H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   462
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   463
apply (erule_tac xa = x in analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   464
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   465
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   466
lemma lemma2: "Key (invKey K) \<in> analz H ==>   
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   467
               insert (Crypt K X) (analz (insert X H)) \<subseteq>  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   468
               analz (insert (Crypt K X) H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   469
apply auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   470
apply (erule_tac xa = x in analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   471
apply (blast intro: analz_insertI analz.Decrypt)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   472
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   473
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   474
lemma analz_insert_Decrypt:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   475
     "Key (invKey K) \<in> analz H ==>   
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   476
               analz (insert (Crypt K X) H) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   477
               insert (Crypt K X) (analz (insert X H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   478
by (intro equalityI lemma1 lemma2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   479
16818
paulson
parents: 16796
diff changeset
   480
text{*Case analysis: either the message is secure, or it is not! Effective,
paulson
parents: 16796
diff changeset
   481
but can cause subgoals to blow up! Use with @{text "split_if"}; apparently
paulson
parents: 16796
diff changeset
   482
@{text "split_tac"} does not cope with patterns such as @{term"analz (insert
paulson
parents: 16796
diff changeset
   483
(Crypt K X) H)"} *} 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   484
lemma analz_Crypt_if [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   485
     "analz (insert (Crypt K X) H) =                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   486
          (if (Key (invKey K) \<in> analz H)                 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   487
           then insert (Crypt K X) (analz (insert X H))  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   488
           else insert (Crypt K X) (analz H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   489
by (simp add: analz_insert_Crypt analz_insert_Decrypt)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   490
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   491
16818
paulson
parents: 16796
diff changeset
   492
text{*This rule supposes "for the sake of argument" that we have the key.*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   493
lemma analz_insert_Crypt_subset:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   494
     "analz (insert (Crypt K X) H) \<subseteq>   
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   495
           insert (Crypt K X) (analz (insert X H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   496
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   497
apply (erule analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   498
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   499
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   500
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   501
lemma analz_image_Key [simp]: "analz (Key`N) = Key`N"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   502
apply auto
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   503
apply (erule analz.induct, auto)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   504
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   505
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   506
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   507
subsubsection{*Idempotence and transitivity *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   508
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   509
lemma analz_analzD [dest!]: "X\<in> analz (analz H) ==> X\<in> analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   510
by (erule analz.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   511
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   512
lemma analz_idem [simp]: "analz (analz H) = analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   513
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   514
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   515
lemma analz_subset_iff [simp]: "(analz G \<subseteq> analz H) = (G \<subseteq> analz H)"
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   516
apply (rule iffI)
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   517
apply (iprover intro: subset_trans analz_increasing)  
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   518
apply (frule analz_mono, simp) 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   519
done
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   520
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   521
lemma analz_trans: "[| X\<in> analz G;  G \<subseteq> analz H |] ==> X\<in> analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   522
by (drule analz_mono, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   523
16818
paulson
parents: 16796
diff changeset
   524
text{*Cut; Lemma 2 of Lowe*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   525
lemma analz_cut: "[| Y\<in> analz (insert X H);  X\<in> analz H |] ==> Y\<in> analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   526
by (erule analz_trans, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   527
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   528
(*Cut can be proved easily by induction on
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   529
   "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   530
*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   531
16818
paulson
parents: 16796
diff changeset
   532
text{*This rewrite rule helps in the simplification of messages that involve
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   533
  the forwarding of unknown components (X).  Without it, removing occurrences
16818
paulson
parents: 16796
diff changeset
   534
  of X can be very complicated. *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   535
lemma analz_insert_eq: "X\<in> analz H ==> analz (insert X H) = analz H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   536
by (blast intro: analz_cut analz_insertI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   537
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   538
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   539
text{*A congruence rule for "analz" *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   540
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   541
lemma analz_subset_cong:
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   542
     "[| analz G \<subseteq> analz G'; analz H \<subseteq> analz H' |] 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   543
      ==> analz (G \<union> H) \<subseteq> analz (G' \<union> H')"
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   544
apply simp
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   545
apply (iprover intro: conjI subset_trans analz_mono Un_upper1 Un_upper2) 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   546
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   547
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   548
lemma analz_cong:
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   549
     "[| analz G = analz G'; analz H = analz H' |] 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   550
      ==> analz (G \<union> H) = analz (G' \<union> H')"
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   551
by (intro equalityI analz_subset_cong, simp_all) 
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   552
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   553
lemma analz_insert_cong:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   554
     "analz H = analz H' ==> analz(insert X H) = analz(insert X H')"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   555
by (force simp only: insert_def intro!: analz_cong)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   556
16818
paulson
parents: 16796
diff changeset
   557
text{*If there are no pairs or encryptions then analz does nothing*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   558
lemma analz_trivial:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   559
     "[| \<forall>X Y. {|X,Y|} \<notin> H;  \<forall>X K. Crypt K X \<notin> H |] ==> analz H = H"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   560
apply safe
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   561
apply (erule analz.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   562
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   563
16818
paulson
parents: 16796
diff changeset
   564
text{*These two are obsolete (with a single Spy) but cost little to prove...*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   565
lemma analz_UN_analz_lemma:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   566
     "X\<in> analz (\<Union>i\<in>A. analz (H i)) ==> X\<in> analz (\<Union>i\<in>A. H i)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   567
apply (erule analz.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   568
apply (blast intro: analz_mono [THEN [2] rev_subsetD])+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   569
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   570
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   571
lemma analz_UN_analz [simp]: "analz (\<Union>i\<in>A. analz (H i)) = analz (\<Union>i\<in>A. H i)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   572
by (blast intro: analz_UN_analz_lemma analz_mono [THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   573
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   574
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   575
subsection{*Inductive relation "synth"*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   576
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   577
text{*Inductive definition of "synth" -- what can be built up from a set of
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   578
    messages.  A form of upward closure.  Pairs can be built, messages
3668
a39baf59ea47 Split base cases from "msg" to "atomic" in order
paulson
parents: 2516
diff changeset
   579
    encrypted with known keys.  Agent names are public domain.
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   580
    Numbers can be guessed, but Nonces cannot be.  *}
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   581
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   582
consts  synth   :: "msg set => msg set"
1913
2809adb15eb0 Renaming of functions, and tidying
paulson
parents: 1839
diff changeset
   583
inductive "synth H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   584
  intros 
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   585
    Inj    [intro]:   "X \<in> H ==> X \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   586
    Agent  [intro]:   "Agent agt \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   587
    Number [intro]:   "Number n  \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   588
    Hash   [intro]:   "X \<in> synth H ==> Hash X \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   589
    MPair  [intro]:   "[|X \<in> synth H;  Y \<in> synth H|] ==> {|X,Y|} \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   590
    Crypt  [intro]:   "[|X \<in> synth H;  Key(K) \<in> H|] ==> Crypt K X \<in> synth H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   591
16818
paulson
parents: 16796
diff changeset
   592
text{*Monotonicity*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   593
lemma synth_mono: "G\<subseteq>H ==> synth(G) \<subseteq> synth(H)"
16818
paulson
parents: 16796
diff changeset
   594
  by (auto, erule synth.induct, auto)  
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   595
16818
paulson
parents: 16796
diff changeset
   596
text{*NO @{text Agent_synth}, as any Agent name can be synthesized.  
paulson
parents: 16796
diff changeset
   597
  The same holds for @{term Number}*}
11192
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   598
inductive_cases Nonce_synth [elim!]: "Nonce n \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   599
inductive_cases Key_synth   [elim!]: "Key K \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   600
inductive_cases Hash_synth  [elim!]: "Hash X \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   601
inductive_cases MPair_synth [elim!]: "{|X,Y|} \<in> synth H"
5fd02b905a9a a few basic X-symbols
paulson
parents: 11189
diff changeset
   602
inductive_cases Crypt_synth [elim!]: "Crypt K X \<in> synth H"
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   603
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   604
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   605
lemma synth_increasing: "H \<subseteq> synth(H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   606
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   607
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   608
subsubsection{*Unions *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   609
16818
paulson
parents: 16796
diff changeset
   610
text{*Converse fails: we can synth more from the union than from the 
paulson
parents: 16796
diff changeset
   611
  separate parts, building a compound message using elements of each.*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   612
lemma synth_Un: "synth(G) \<union> synth(H) \<subseteq> synth(G \<union> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   613
by (intro Un_least synth_mono Un_upper1 Un_upper2)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   614
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   615
lemma synth_insert: "insert X (synth H) \<subseteq> synth(insert X H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   616
by (blast intro: synth_mono [THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   617
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   618
subsubsection{*Idempotence and transitivity *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   619
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   620
lemma synth_synthD [dest!]: "X\<in> synth (synth H) ==> X\<in> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   621
by (erule synth.induct, blast+)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   622
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   623
lemma synth_idem: "synth (synth H) = synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   624
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   625
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   626
lemma synth_subset_iff [simp]: "(synth G \<subseteq> synth H) = (G \<subseteq> synth H)"
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   627
apply (rule iffI)
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   628
apply (iprover intro: subset_trans synth_increasing)  
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   629
apply (frule synth_mono, simp add: synth_idem) 
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   630
done
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   631
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   632
lemma synth_trans: "[| X\<in> synth G;  G \<subseteq> synth H |] ==> X\<in> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   633
by (drule synth_mono, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   634
16818
paulson
parents: 16796
diff changeset
   635
text{*Cut; Lemma 2 of Lowe*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   636
lemma synth_cut: "[| Y\<in> synth (insert X H);  X\<in> synth H |] ==> Y\<in> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   637
by (erule synth_trans, blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   638
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   639
lemma Agent_synth [simp]: "Agent A \<in> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   640
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   641
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   642
lemma Number_synth [simp]: "Number n \<in> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   643
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   644
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   645
lemma Nonce_synth_eq [simp]: "(Nonce N \<in> synth H) = (Nonce N \<in> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   646
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   647
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   648
lemma Key_synth_eq [simp]: "(Key K \<in> synth H) = (Key K \<in> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   649
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   650
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   651
lemma Crypt_synth_eq [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   652
     "Key K \<notin> H ==> (Crypt K X \<in> synth H) = (Crypt K X \<in> H)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   653
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   654
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   655
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   656
lemma keysFor_synth [simp]: 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   657
    "keysFor (synth H) = keysFor H \<union> invKey`{K. Key K \<in> H}"
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   658
by (unfold keysFor_def, blast)
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   659
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   660
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   661
subsubsection{*Combinations of parts, analz and synth *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   662
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   663
lemma parts_synth [simp]: "parts (synth H) = parts H \<union> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   664
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   665
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   666
apply (erule parts.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   667
apply (blast intro: synth_increasing [THEN parts_mono, THEN subsetD] 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   668
                    parts.Fst parts.Snd parts.Body)+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   669
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   670
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   671
lemma analz_analz_Un [simp]: "analz (analz G \<union> H) = analz (G \<union> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   672
apply (intro equalityI analz_subset_cong)+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   673
apply simp_all
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   674
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   675
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   676
lemma analz_synth_Un [simp]: "analz (synth G \<union> H) = analz (G \<union> H) \<union> synth G"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   677
apply (rule equalityI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   678
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   679
apply (erule analz.induct)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   680
prefer 5 apply (blast intro: analz_mono [THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   681
apply (blast intro: analz.Fst analz.Snd analz.Decrypt)+
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   682
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   683
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   684
lemma analz_synth [simp]: "analz (synth H) = analz H \<union> synth H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   685
apply (cut_tac H = "{}" in analz_synth_Un)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   686
apply (simp (no_asm_use))
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   687
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   688
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   689
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   690
subsubsection{*For reasoning about the Fake rule in traces *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   691
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   692
lemma parts_insert_subset_Un: "X\<in> G ==> parts(insert X H) \<subseteq> parts G \<union> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   693
by (rule subset_trans [OF parts_mono parts_Un_subset2], blast)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   694
16818
paulson
parents: 16796
diff changeset
   695
text{*More specifically for Fake.  Very occasionally we could do with a version
paulson
parents: 16796
diff changeset
   696
  of the form  @{term"parts{X} \<subseteq> synth (analz H) \<union> parts H"} *}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   697
lemma Fake_parts_insert:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   698
     "X \<in> synth (analz H) ==>  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   699
      parts (insert X H) \<subseteq> synth (analz H) \<union> parts H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   700
apply (drule parts_insert_subset_Un)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   701
apply (simp (no_asm_use))
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   702
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   703
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   704
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   705
lemma Fake_parts_insert_in_Un:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   706
     "[|Z \<in> parts (insert X H);  X: synth (analz H)|] 
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   707
      ==> Z \<in>  synth (analz H) \<union> parts H";
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   708
by (blast dest: Fake_parts_insert  [THEN subsetD, dest])
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   709
16818
paulson
parents: 16796
diff changeset
   710
text{*@{term H} is sometimes @{term"Key ` KK \<union> spies evs"}, so can't put 
paulson
parents: 16796
diff changeset
   711
  @{term "G=H"}.*}
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   712
lemma Fake_analz_insert:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   713
     "X\<in> synth (analz G) ==>  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   714
      analz (insert X H) \<subseteq> synth (analz G) \<union> analz (G \<union> H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   715
apply (rule subsetI)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   716
apply (subgoal_tac "x \<in> analz (synth (analz G) \<union> H) ")
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   717
prefer 2 apply (blast intro: analz_mono [THEN [2] rev_subsetD] analz_mono [THEN synth_mono, THEN [2] rev_subsetD])
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   718
apply (simp (no_asm_use))
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   719
apply blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   720
done
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   721
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   722
lemma analz_conj_parts [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   723
     "(X \<in> analz H & X \<in> parts H) = (X \<in> analz H)"
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 14126
diff changeset
   724
by (blast intro: analz_subset_parts [THEN subsetD])
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   725
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   726
lemma analz_disj_parts [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   727
     "(X \<in> analz H | X \<in> parts H) = (X \<in> parts H)"
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 14126
diff changeset
   728
by (blast intro: analz_subset_parts [THEN subsetD])
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   729
16818
paulson
parents: 16796
diff changeset
   730
text{*Without this equation, other rules for synth and analz would yield
paulson
parents: 16796
diff changeset
   731
  redundant cases*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   732
lemma MPair_synth_analz [iff]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   733
     "({|X,Y|} \<in> synth (analz H)) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   734
      (X \<in> synth (analz H) & Y \<in> synth (analz H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   735
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   736
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   737
lemma Crypt_synth_analz:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   738
     "[| Key K \<in> analz H;  Key (invKey K) \<in> analz H |]  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   739
       ==> (Crypt K X \<in> synth (analz H)) = (X \<in> synth (analz H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   740
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   741
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   742
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   743
lemma Hash_synth_analz [simp]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   744
     "X \<notin> synth (analz H)  
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   745
      ==> (Hash{|X,Y|} \<in> synth (analz H)) = (Hash{|X,Y|} \<in> analz H)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   746
by blast
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   747
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   748
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   749
subsection{*HPair: a combination of Hash and MPair*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   750
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   751
subsubsection{*Freeness *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   752
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   753
lemma Agent_neq_HPair: "Agent A ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   754
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   755
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   756
lemma Nonce_neq_HPair: "Nonce N ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   757
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   758
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   759
lemma Number_neq_HPair: "Number N ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   760
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   761
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   762
lemma Key_neq_HPair: "Key K ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   763
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   764
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   765
lemma Hash_neq_HPair: "Hash Z ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   766
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   767
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   768
lemma Crypt_neq_HPair: "Crypt K X' ~= Hash[X] Y"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   769
by (unfold HPair_def, simp)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   770
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   771
lemmas HPair_neqs = Agent_neq_HPair Nonce_neq_HPair Number_neq_HPair 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   772
                    Key_neq_HPair Hash_neq_HPair Crypt_neq_HPair
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   773
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   774
declare HPair_neqs [iff]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   775
declare HPair_neqs [symmetric, iff]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   776
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   777
lemma HPair_eq [iff]: "(Hash[X'] Y' = Hash[X] Y) = (X' = X & Y'=Y)"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   778
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   779
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   780
lemma MPair_eq_HPair [iff]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   781
     "({|X',Y'|} = Hash[X] Y) = (X' = Hash{|X,Y|} & Y'=Y)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   782
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   783
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   784
lemma HPair_eq_MPair [iff]:
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   785
     "(Hash[X] Y = {|X',Y'|}) = (X' = Hash{|X,Y|} & Y'=Y)"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   786
by (auto simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   787
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   788
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   789
subsubsection{*Specialized laws, proved in terms of those for Hash and MPair *}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   790
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   791
lemma keysFor_insert_HPair [simp]: "keysFor (insert (Hash[X] Y) H) = keysFor H"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   792
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   793
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   794
lemma parts_insert_HPair [simp]: 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   795
    "parts (insert (Hash[X] Y) H) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   796
     insert (Hash[X] Y) (insert (Hash{|X,Y|}) (parts (insert Y H)))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   797
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   798
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   799
lemma analz_insert_HPair [simp]: 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   800
    "analz (insert (Hash[X] Y) H) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   801
     insert (Hash[X] Y) (insert (Hash{|X,Y|}) (analz (insert Y H)))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   802
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   803
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   804
lemma HPair_synth_analz [simp]:
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   805
     "X \<notin> synth (analz H)  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   806
    ==> (Hash[X] Y \<in> synth (analz H)) =  
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   807
        (Hash {|X, Y|} \<in> analz H & Y \<in> synth (analz H))"
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   808
by (simp add: HPair_def)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   809
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   810
16818
paulson
parents: 16796
diff changeset
   811
text{*We do NOT want Crypt... messages broken up in protocols!!*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   812
declare parts.Body [rule del]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   813
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   814
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   815
text{*Rewrites to push in Key and Crypt messages, so that other messages can
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   816
    be pulled out using the @{text analz_insert} rules*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   817
ML
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   818
{*
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   819
fun insComm x y = inst "x" x (inst "y" y insert_commute);
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   820
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   821
bind_thms ("pushKeys",
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   822
           map (insComm "Key ?K") 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   823
                   ["Agent ?C", "Nonce ?N", "Number ?N", 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   824
		    "Hash ?X", "MPair ?X ?Y", "Crypt ?X ?K'"]);
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   825
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   826
bind_thms ("pushCrypts",
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   827
           map (insComm "Crypt ?X ?K") 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   828
                     ["Agent ?C", "Nonce ?N", "Number ?N", 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   829
		      "Hash ?X'", "MPair ?X' ?Y"]);
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   830
*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   831
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   832
text{*Cannot be added with @{text "[simp]"} -- messages should not always be
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   833
  re-ordered. *}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   834
lemmas pushes = pushKeys pushCrypts
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   835
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   836
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   837
subsection{*Tactics useful for many protocol proofs*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   838
ML
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   839
{*
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   840
val invKey = thm "invKey"
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   841
val keysFor_def = thm "keysFor_def"
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   842
val HPair_def = thm "HPair_def"
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   843
val symKeys_def = thm "symKeys_def"
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   844
val parts_mono = thm "parts_mono";
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   845
val analz_mono = thm "analz_mono";
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   846
val synth_mono = thm "synth_mono";
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   847
val analz_increasing = thm "analz_increasing";
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   848
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   849
val analz_insertI = thm "analz_insertI";
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   850
val analz_subset_parts = thm "analz_subset_parts";
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   851
val Fake_parts_insert = thm "Fake_parts_insert";
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   852
val Fake_analz_insert = thm "Fake_analz_insert";
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   853
val pushes = thms "pushes";
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   854
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   855
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   856
(*Prove base case (subgoal i) and simplify others.  A typical base case
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   857
  concerns  Crypt K X \<notin> Key`shrK`bad  and cannot be proved by rewriting
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   858
  alone.*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   859
fun prove_simple_subgoals_tac i = 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   860
    force_tac (claset(), simpset() addsimps [image_eq_UN]) i THEN
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   861
    ALLGOALS Asm_simp_tac
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   862
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   863
(*Analysis of Fake cases.  Also works for messages that forward unknown parts,
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   864
  but this application is no longer necessary if analz_insert_eq is used.
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   865
  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   866
  DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   867
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   868
(*Apply rules to break down assumptions of the form
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   869
  Y \<in> parts(insert X H)  and  Y \<in> analz(insert X H)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   870
*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   871
val Fake_insert_tac = 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   872
    dresolve_tac [impOfSubs Fake_analz_insert,
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   873
                  impOfSubs Fake_parts_insert] THEN'
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   874
    eresolve_tac [asm_rl, thm"synth.Inj"];
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   875
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   876
fun Fake_insert_simp_tac ss i = 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   877
    REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i;
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   878
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   879
fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   880
    (Fake_insert_simp_tac ss 1
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   881
     THEN
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   882
     IF_UNSOLVED (Blast.depth_tac
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   883
		  (cs addIs [analz_insertI,
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   884
				   impOfSubs analz_subset_parts]) 4 1))
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   885
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   886
(*The explicit claset and simpset arguments help it work with Isar*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   887
fun gen_spy_analz_tac (cs,ss) i =
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   888
  DETERM
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   889
   (SELECT_GOAL
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   890
     (EVERY 
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   891
      [  (*push in occurrences of X...*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   892
       (REPEAT o CHANGED)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   893
           (res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   894
       (*...allowing further simplifications*)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   895
       simp_tac ss 1,
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   896
       REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   897
       DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i)
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   898
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   899
fun spy_analz_tac i = gen_spy_analz_tac (claset(), simpset()) i
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   900
*}
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   901
16818
paulson
parents: 16796
diff changeset
   902
text{*By default only @{text o_apply} is built-in.  But in the presence of
paulson
parents: 16796
diff changeset
   903
eta-expansion this means that some terms displayed as @{term "f o g"} will be
paulson
parents: 16796
diff changeset
   904
rewritten, and others will not!*}
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   905
declare o_def [simp]
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   906
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   907
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   908
lemma Crypt_notin_image_Key [simp]: "Crypt K X \<notin> Key ` A"
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   909
by auto
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   910
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   911
lemma Hash_notin_image_Key [simp] :"Hash X \<notin> Key ` A"
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   912
by auto
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   913
14200
d8598e24f8fa Removal of the Key_supply axiom (affects many possbility proofs) and minor
paulson
parents: 14181
diff changeset
   914
lemma synth_analz_mono: "G\<subseteq>H ==> synth (analz(G)) \<subseteq> synth (analz(H))"
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   915
by (iprover intro: synth_mono analz_mono) 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   916
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   917
lemma Fake_analz_eq [simp]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   918
     "X \<in> synth(analz H) ==> synth (analz (insert X H)) = synth (analz H)"
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   919
apply (drule Fake_analz_insert[of _ _ "H"])
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   920
apply (simp add: synth_increasing[THEN Un_absorb2])
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   921
apply (drule synth_mono)
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   922
apply (simp add: synth_idem)
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   923
apply (rule equalityI)
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   924
apply (simp add: );
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   925
apply (rule synth_analz_mono, blast)   
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   926
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   927
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   928
text{*Two generalizations of @{text analz_insert_eq}*}
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   929
lemma gen_analz_insert_eq [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   930
     "X \<in> analz H ==> ALL G. H \<subseteq> G --> analz (insert X G) = analz G";
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   931
by (blast intro: analz_cut analz_insertI analz_mono [THEN [2] rev_subsetD])
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   932
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   933
lemma synth_analz_insert_eq [rule_format]:
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   934
     "X \<in> synth (analz H) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   935
      ==> ALL G. H \<subseteq> G --> (Key K \<in> analz (insert X G)) = (Key K \<in> analz G)";
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   936
apply (erule synth.induct) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   937
apply (simp_all add: gen_analz_insert_eq subset_trans [OF _ subset_insertI]) 
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   938
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   939
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   940
lemma Fake_parts_sing:
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   941
     "X \<in> synth (analz H) ==> parts{X} \<subseteq> synth (analz H) \<union> parts H";
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   942
apply (rule subset_trans) 
17689
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   943
 apply (erule_tac [2] Fake_parts_insert)
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   944
apply (rule parts_mono)  
a04b5b43625e streamlined theory; conformance to recent publication
paulson
parents: 16818
diff changeset
   945
apply (blast intro: elim:); 
13922
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   946
done
75ae4244a596 Changes required by the certified email protocol
paulson
parents: 11270
diff changeset
   947
14145
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 14126
diff changeset
   948
lemmas Fake_parts_sing_imp_Un = Fake_parts_sing [THEN [2] rev_subsetD]
2e31b8cc8788 ZhouGollmann: new example (fair non-repudiation protocol)
paulson
parents: 14126
diff changeset
   949
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   950
method_setup spy_analz = {*
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   951
    Method.ctxt_args (fn ctxt =>
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   952
        Method.METHOD (fn facts => 
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14200
diff changeset
   953
            gen_spy_analz_tac (local_clasimpset_of ctxt) 1)) *}
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents: 10833
diff changeset
   954
    "for proving the Fake case when analz is involved"
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   955
11264
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   956
method_setup atomic_spy_analz = {*
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   957
    Method.ctxt_args (fn ctxt =>
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   958
        Method.METHOD (fn facts => 
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14200
diff changeset
   959
            atomic_spy_analz_tac (local_clasimpset_of ctxt) 1)) *}
11264
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   960
    "for debugging spy_analz"
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   961
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   962
method_setup Fake_insert_simp = {*
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   963
    Method.ctxt_args (fn ctxt =>
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   964
        Method.METHOD (fn facts =>
15032
02aed07e01bf local_cla/simpset_of;
wenzelm
parents: 14200
diff changeset
   965
            Fake_insert_simp_tac (local_simpset_of ctxt) 1)) *}
11264
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   966
    "for debugging spy_analz"
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   967
13926
6e62e5357a10 converting more HOL-Auth to new-style theories
paulson
parents: 13922
diff changeset
   968
1839
199243afac2b Proving safety properties of authentication protocols
paulson
parents:
diff changeset
   969
end