src/HOL/Auth/Message_lemmas.ML
author nipkow
Mon, 06 Aug 2001 13:43:24 +0200
changeset 11464 ddea204de5bc
parent 11270 a315a3862bb4
permissions -rw-r--r--
turned translation for 1::nat into def. introduced 1' and replaced most occurrences of 1 by 1'.
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     1
(*  Title:      HOL/Auth/Message
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     2
    ID:         $Id$
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     4
    Copyright   1996  University of Cambridge
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     5
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     6
Datatypes of agents and messages;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     7
Inductive relations "parts", "analz" and "synth"
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     8
*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
     9
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    10
(*ML bindings for definitions and axioms*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    11
val invKey = thm "invKey";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    12
val keysFor_def = thm "keysFor_def";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    13
val parts_mono = thm "parts_mono";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    14
val analz_mono = thm "analz_mono";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    15
val synth_mono = thm "synth_mono";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    16
val HPair_def = thm "HPair_def";
11230
756c5034f08b misc tidying; changing the predicate isSymKey to the set symKeys
paulson
parents: 11217
diff changeset
    17
val symKeys_def = thm "symKeys_def";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    18
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    19
structure parts =
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    20
  struct
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    21
  val induct = thm "parts.induct";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    22
  val Inj    = thm "parts.Inj";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    23
  val Fst    = thm "parts.Fst";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    24
  val Snd    = thm "parts.Snd";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    25
  val Body   = thm "parts.Body";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    26
  end;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    27
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    28
structure analz =
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    29
  struct
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    30
  val induct = thm "analz.induct";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    31
  val Inj    = thm "analz.Inj";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    32
  val Fst    = thm "analz.Fst";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    33
  val Snd    = thm "analz.Snd";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    34
  val Decrypt = thm "analz.Decrypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    35
  end;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    36
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    37
structure synth =
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    38
  struct
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    39
  val induct = thm "synth.induct";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    40
  val Inj    = thm "synth.Inj";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    41
  val Agent  = thm "synth.Agent";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    42
  val Number = thm "synth.Number";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    43
  val Hash   = thm "synth.Hash";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    44
  val Crypt  = thm "synth.Crypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    45
  end;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    46
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    47
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    48
(*Equations hold because constructors are injective; cannot prove for all f*)
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
    49
Goal "(Friend x \\<in> Friend`A) = (x:A)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    50
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    51
qed "Friend_image_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    52
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    53
Goal "(Key x \\<in> Key`A) = (x\\<in>A)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    54
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    55
qed "Key_image_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    56
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
    57
Goal "(Nonce x \\<notin> Key`A)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    58
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    59
qed "Nonce_Key_image_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    60
Addsimps [Friend_image_eq, Key_image_eq, Nonce_Key_image_eq];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    61
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    62
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    63
(** Inverse of keys **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    64
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    65
Goal "(invKey K = invKey K') = (K=K')";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    66
by Safe_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    67
by (rtac box_equals 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    68
by (REPEAT (rtac invKey 2));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    69
by (Asm_simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    70
qed "invKey_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    71
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    72
Addsimps [invKey_eq];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    73
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    74
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    75
(**** keysFor operator ****)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    76
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    77
Goalw [keysFor_def] "keysFor {} = {}";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    78
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    79
qed "keysFor_empty";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    80
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    81
Goalw [keysFor_def] "keysFor (H Un H') = keysFor H Un keysFor H'";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    82
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    83
qed "keysFor_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    84
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    85
Goalw [keysFor_def] "keysFor (\\<Union>i\\<in>A. H i) = (\\<Union>i\\<in>A. keysFor (H i))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    86
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    87
qed "keysFor_UN";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    88
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    89
(*Monotonicity*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
    90
Goalw [keysFor_def] "G\\<subseteq>H ==> keysFor(G) \\<subseteq> keysFor(H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    91
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    92
qed "keysFor_mono";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    93
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    94
Goalw [keysFor_def] "keysFor (insert (Agent A) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    95
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    96
qed "keysFor_insert_Agent";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    97
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    98
Goalw [keysFor_def] "keysFor (insert (Nonce N) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
    99
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   100
qed "keysFor_insert_Nonce";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   101
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   102
Goalw [keysFor_def] "keysFor (insert (Number N) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   103
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   104
qed "keysFor_insert_Number";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   105
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   106
Goalw [keysFor_def] "keysFor (insert (Key K) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   107
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   108
qed "keysFor_insert_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   109
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   110
Goalw [keysFor_def] "keysFor (insert (Hash X) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   111
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   112
qed "keysFor_insert_Hash";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   113
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   114
Goalw [keysFor_def] "keysFor (insert {|X,Y|} H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   115
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   116
qed "keysFor_insert_MPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   117
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   118
Goalw [keysFor_def]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   119
    "keysFor (insert (Crypt K X) H) = insert (invKey K) (keysFor H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   120
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   121
qed "keysFor_insert_Crypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   122
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   123
Addsimps [keysFor_empty, keysFor_Un, keysFor_UN, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   124
          keysFor_insert_Agent, keysFor_insert_Nonce, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   125
	  keysFor_insert_Number, keysFor_insert_Key, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   126
          keysFor_insert_Hash, keysFor_insert_MPair, keysFor_insert_Crypt];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   127
AddSEs [keysFor_Un RS equalityD1 RS subsetD RS UnE,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   128
	keysFor_UN RS equalityD1 RS subsetD RS UN_E];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   129
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   130
Goalw [keysFor_def] "keysFor (Key`E) = {}";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   131
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   132
qed "keysFor_image_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   133
Addsimps [keysFor_image_Key];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   134
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   135
Goalw [keysFor_def] "Crypt K X \\<in> H ==> invKey K \\<in> keysFor H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   136
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   137
qed "Crypt_imp_invKey_keysFor";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   138
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   139
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   140
(**** Inductive relation "parts" ****)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   141
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   142
val major::prems = 
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   143
Goal "[| {|X,Y|} \\<in> parts H;       \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   144
\         [| X \\<in> parts H; Y \\<in> parts H |] ==> P  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   145
\     |] ==> P";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   146
by (cut_facts_tac [major] 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   147
by (resolve_tac prems 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   148
by (REPEAT (eresolve_tac [asm_rl, parts.Fst, parts.Snd] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   149
qed "MPair_parts";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   150
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   151
AddSEs [MPair_parts, make_elim parts.Body];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   152
(*NB These two rules are UNSAFE in the formal sense, as they discard the
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   153
     compound message.  They work well on THIS FILE.  
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   154
  MPair_parts is left as SAFE because it speeds up proofs.
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   155
  The Crypt rule is normally kept UNSAFE to avoid breaking up certificates.*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   156
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   157
Goal "H \\<subseteq> parts(H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   158
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   159
qed "parts_increasing";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   160
11264
a47a9288f3f6 (rough) conversion of Auth/Recur to Isar format
paulson
parents: 11251
diff changeset
   161
bind_thm ("parts_insertI", impOfSubs (subset_insertI RS parts_mono));
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   162
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   163
Goal "parts{} = {}";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   164
by Safe_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   165
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   166
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   167
qed "parts_empty";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   168
Addsimps [parts_empty];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   169
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   170
Goal "X\\<in> parts{} ==> P";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   171
by (Asm_full_simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   172
qed "parts_emptyE";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   173
AddSEs [parts_emptyE];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   174
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   175
(*WARNING: loops if H = {Y}, therefore must not be repeated!*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   176
Goal "X\\<in> parts H ==> \\<exists>Y\\<in>H. X\\<in> parts {Y}";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   177
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   178
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   179
qed "parts_singleton";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   180
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   181
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   182
(** Unions **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   183
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   184
Goal "parts(G) Un parts(H) \\<subseteq> parts(G Un H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   185
by (REPEAT (ares_tac [Un_least, parts_mono, Un_upper1, Un_upper2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   186
val parts_Un_subset1 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   187
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   188
Goal "parts(G Un H) \\<subseteq> parts(G) Un parts(H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   189
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   190
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   191
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   192
val parts_Un_subset2 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   193
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   194
Goal "parts(G Un H) = parts(G) Un parts(H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   195
by (REPEAT (ares_tac [equalityI, parts_Un_subset1, parts_Un_subset2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   196
qed "parts_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   197
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   198
Goal "parts (insert X H) = parts {X} Un parts H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   199
by (stac (read_instantiate [("A","H")] insert_is_Un) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   200
by (simp_tac (HOL_ss addsimps [parts_Un]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   201
qed "parts_insert";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   202
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   203
(*TWO inserts to avoid looping.  This rewrite is better than nothing.
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   204
  Not suitable for Addsimps: its behaviour can be strange.*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   205
Goal "parts (insert X (insert Y H)) = parts {X} Un parts {Y} Un parts H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   206
by (simp_tac (simpset() addsimps [Un_assoc]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   207
by (simp_tac (simpset() addsimps [parts_insert RS sym]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   208
qed "parts_insert2";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   209
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   210
Goal "(\\<Union>x\\<in>A. parts(H x)) \\<subseteq> parts(\\<Union>x\\<in>A. H x)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   211
by (REPEAT (ares_tac [UN_least, parts_mono, UN_upper] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   212
val parts_UN_subset1 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   213
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   214
Goal "parts(\\<Union>x\\<in>A. H x) \\<subseteq> (\\<Union>x\\<in>A. parts(H x))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   215
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   216
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   217
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   218
val parts_UN_subset2 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   219
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   220
Goal "parts(\\<Union>x\\<in>A. H x) = (\\<Union>x\\<in>A. parts(H x))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   221
by (REPEAT (ares_tac [equalityI, parts_UN_subset1, parts_UN_subset2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   222
qed "parts_UN";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   223
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   224
(*Added to simplify arguments to parts, analz and synth.
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   225
  NOTE: the UN versions are no longer used!*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   226
Addsimps [parts_Un, parts_UN];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   227
AddSEs [parts_Un RS equalityD1 RS subsetD RS UnE,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   228
	parts_UN RS equalityD1 RS subsetD RS UN_E];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   229
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   230
Goal "insert X (parts H) \\<subseteq> parts(insert X H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   231
by (blast_tac (claset() addIs [impOfSubs parts_mono]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   232
qed "parts_insert_subset";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   233
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   234
(** Idempotence and transitivity **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   235
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   236
Goal "X\\<in> parts (parts H) ==> X\\<in> parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   237
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   238
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   239
qed "parts_partsD";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   240
AddSDs [parts_partsD];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   241
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   242
Goal "parts (parts H) = parts H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   243
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   244
qed "parts_idem";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   245
Addsimps [parts_idem];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   246
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   247
Goal "[| X\\<in> parts G;  G \\<subseteq> parts H |] ==> X\\<in> parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   248
by (dtac parts_mono 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   249
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   250
qed "parts_trans";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   251
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   252
(*Cut*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   253
Goal "[| Y\\<in> parts (insert X G);  X\\<in> parts H |] \
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   254
\              ==> Y\\<in> parts (G Un H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   255
by (etac parts_trans 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   256
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   257
qed "parts_cut";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   258
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   259
Goal "X\\<in> parts H ==> parts (insert X H) = parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   260
by (fast_tac (claset() addSDs [parts_cut]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   261
                       addIs  [parts_insertI] 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   262
                       addss (simpset())) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   263
qed "parts_cut_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   264
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   265
Addsimps [parts_cut_eq];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   266
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   267
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   268
(** Rewrite rules for pulling out atomic messages **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   269
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   270
fun parts_tac i =
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   271
  EVERY [rtac ([subsetI, parts_insert_subset] MRS equalityI) i,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   272
         etac parts.induct i,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   273
         Auto_tac];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   274
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   275
Goal "parts (insert (Agent agt) H) = insert (Agent agt) (parts H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   276
by (parts_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   277
qed "parts_insert_Agent";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   278
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   279
Goal "parts (insert (Nonce N) H) = insert (Nonce N) (parts H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   280
by (parts_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   281
qed "parts_insert_Nonce";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   282
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   283
Goal "parts (insert (Number N) H) = insert (Number N) (parts H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   284
by (parts_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   285
qed "parts_insert_Number";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   286
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   287
Goal "parts (insert (Key K) H) = insert (Key K) (parts H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   288
by (parts_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   289
qed "parts_insert_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   290
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   291
Goal "parts (insert (Hash X) H) = insert (Hash X) (parts H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   292
by (parts_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   293
qed "parts_insert_Hash";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   294
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   295
Goal "parts (insert (Crypt K X) H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   296
\         insert (Crypt K X) (parts (insert X H))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   297
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   298
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   299
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   300
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   301
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   302
by (ALLGOALS (blast_tac (claset() addIs [parts.Body])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   303
qed "parts_insert_Crypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   304
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   305
Goal "parts (insert {|X,Y|} H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   306
\         insert {|X,Y|} (parts (insert X (insert Y H)))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   307
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   308
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   309
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   310
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   311
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   312
by (ALLGOALS (blast_tac (claset() addIs [parts.Fst, parts.Snd])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   313
qed "parts_insert_MPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   314
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   315
Addsimps [parts_insert_Agent, parts_insert_Nonce, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   316
	  parts_insert_Number, parts_insert_Key, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   317
          parts_insert_Hash, parts_insert_Crypt, parts_insert_MPair];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   318
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   319
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   320
Goal "parts (Key`N) = Key`N";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   321
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   322
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   323
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   324
qed "parts_image_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   325
Addsimps [parts_image_Key];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   326
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   327
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   328
(*In any message, there is an upper bound N on its greatest nonce.*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   329
Goal "\\<exists>N. \\<forall>n. N\\<le>n --> Nonce n \\<notin> parts {msg}";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   330
by (induct_tac "msg" 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   331
by (ALLGOALS (asm_simp_tac (simpset() addsimps [exI, parts_insert2])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   332
(*MPair case: blast_tac works out the necessary sum itself!*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   333
by (blast_tac (claset() addSEs [add_leE]) 2);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   334
(*Nonce case*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   335
by (res_inst_tac [("x","N + Suc nat")] exI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   336
by (auto_tac (claset() addSEs [add_leE], simpset()));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   337
qed "msg_Nonce_supply";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   338
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   339
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   340
(**** Inductive relation "analz" ****)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   341
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   342
val major::prems = 
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   343
Goal "[| {|X,Y|} \\<in> analz H;       \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   344
\            [| X \\<in> analz H; Y \\<in> analz H |] ==> P  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   345
\         |] ==> P";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   346
by (cut_facts_tac [major] 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   347
by (resolve_tac prems 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   348
by (REPEAT (eresolve_tac [asm_rl, analz.Fst, analz.Snd] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   349
qed "MPair_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   350
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   351
AddSEs [MPair_analz];     (*Making it safe speeds up proofs*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   352
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   353
Goal "H \\<subseteq> analz(H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   354
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   355
qed "analz_increasing";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   356
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   357
Goal "analz H \\<subseteq> parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   358
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   359
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   360
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   361
qed "analz_subset_parts";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   362
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   363
bind_thm ("not_parts_not_analz", analz_subset_parts RS contra_subsetD);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   364
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   365
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   366
Goal "parts (analz H) = parts H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   367
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   368
by (rtac (analz_subset_parts RS parts_mono RS subset_trans) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   369
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   370
by (blast_tac (claset() addIs [analz_increasing RS parts_mono RS subsetD]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   371
qed "parts_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   372
Addsimps [parts_analz];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   373
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   374
Goal "analz (parts H) = parts H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   375
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   376
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   377
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   378
qed "analz_parts";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   379
Addsimps [analz_parts];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   380
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   381
bind_thm ("analz_insertI", impOfSubs (subset_insertI RS analz_mono));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   382
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   383
(** General equational properties **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   384
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   385
Goal "analz{} = {}";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   386
by Safe_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   387
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   388
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   389
qed "analz_empty";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   390
Addsimps [analz_empty];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   391
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   392
(*Converse fails: we can analz more from the union than from the 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   393
  separate parts, as a key in one might decrypt a message in the other*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   394
Goal "analz(G) Un analz(H) \\<subseteq> analz(G Un H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   395
by (REPEAT (ares_tac [Un_least, analz_mono, Un_upper1, Un_upper2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   396
qed "analz_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   397
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   398
Goal "insert X (analz H) \\<subseteq> analz(insert X H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   399
by (blast_tac (claset() addIs [impOfSubs analz_mono]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   400
qed "analz_insert";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   401
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   402
(** Rewrite rules for pulling out atomic messages **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   403
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   404
fun analz_tac i =
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   405
  EVERY [rtac ([subsetI, analz_insert] MRS equalityI) i,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   406
         etac analz.induct i,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   407
         Auto_tac];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   408
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   409
Goal "analz (insert (Agent agt) H) = insert (Agent agt) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   410
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   411
qed "analz_insert_Agent";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   412
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   413
Goal "analz (insert (Nonce N) H) = insert (Nonce N) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   414
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   415
qed "analz_insert_Nonce";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   416
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   417
Goal "analz (insert (Number N) H) = insert (Number N) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   418
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   419
qed "analz_insert_Number";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   420
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   421
Goal "analz (insert (Hash X) H) = insert (Hash X) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   422
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   423
qed "analz_insert_Hash";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   424
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   425
(*Can only pull out Keys if they are not needed to decrypt the rest*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   426
Goalw [keysFor_def]
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   427
    "K \\<notin> keysFor (analz H) ==>  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   428
\         analz (insert (Key K) H) = insert (Key K) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   429
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   430
qed "analz_insert_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   431
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   432
Goal "analz (insert {|X,Y|} H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   433
\         insert {|X,Y|} (analz (insert X (insert Y H)))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   434
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   435
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   436
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   437
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   438
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   439
by (ALLGOALS (blast_tac (claset() addIs [analz.Fst, analz.Snd])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   440
qed "analz_insert_MPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   441
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   442
(*Can pull out enCrypted message if the Key is not known*)
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   443
Goal "Key (invKey K) \\<notin> analz H ==>  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   444
\              analz (insert (Crypt K X) H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   445
\              insert (Crypt K X) (analz H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   446
by (analz_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   447
qed "analz_insert_Crypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   448
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   449
Goal "Key (invKey K) \\<in> analz H ==>  \
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   450
\              analz (insert (Crypt K X) H) \\<subseteq> \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   451
\              insert (Crypt K X) (analz (insert X H))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   452
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   453
by (eres_inst_tac [("xa","x")] analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   454
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   455
val lemma1 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   456
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   457
Goal "Key (invKey K) \\<in> analz H ==>  \
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   458
\              insert (Crypt K X) (analz (insert X H)) \\<subseteq> \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   459
\              analz (insert (Crypt K X) H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   460
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   461
by (eres_inst_tac [("xa","x")] analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   462
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   463
by (blast_tac (claset() addIs [analz_insertI, analz.Decrypt]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   464
val lemma2 = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   465
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   466
Goal "Key (invKey K) \\<in> analz H ==>  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   467
\              analz (insert (Crypt K X) H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   468
\              insert (Crypt K X) (analz (insert X H))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   469
by (REPEAT (ares_tac [equalityI, lemma1, lemma2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   470
qed "analz_insert_Decrypt";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   471
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   472
(*Case analysis: either the message is secure, or it is not!
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   473
  Effective, but can cause subgoals to blow up!
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   474
  Use with split_if;  apparently split_tac does not cope with patterns
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   475
  such as "analz (insert (Crypt K X) H)" *)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   476
Goal "analz (insert (Crypt K X) H) =                \
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   477
\         (if (Key (invKey K) \\<in> analz H)                \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   478
\          then insert (Crypt K X) (analz (insert X H)) \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   479
\          else insert (Crypt K X) (analz H))";
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   480
by (case_tac "Key (invKey K)  \\<in> analz H " 1);
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   481
by (ALLGOALS (asm_simp_tac (simpset() addsimps [analz_insert_Crypt, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   482
						analz_insert_Decrypt])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   483
qed "analz_Crypt_if";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   484
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   485
Addsimps [analz_insert_Agent, analz_insert_Nonce, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   486
	  analz_insert_Number, analz_insert_Key, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   487
          analz_insert_Hash, analz_insert_MPair, analz_Crypt_if];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   488
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   489
(*This rule supposes "for the sake of argument" that we have the key.*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   490
Goal  "analz (insert (Crypt K X) H) \\<subseteq>  \
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   491
\          insert (Crypt K X) (analz (insert X H))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   492
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   493
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   494
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   495
qed "analz_insert_Crypt_subset";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   496
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   497
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   498
Goal "analz (Key`N) = Key`N";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   499
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   500
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   501
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   502
qed "analz_image_Key";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   503
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   504
Addsimps [analz_image_Key];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   505
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   506
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   507
(** Idempotence and transitivity **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   508
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   509
Goal "X\\<in> analz (analz H) ==> X\\<in> analz H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   510
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   511
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   512
qed "analz_analzD";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   513
AddSDs [analz_analzD];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   514
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   515
Goal "analz (analz H) = analz H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   516
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   517
qed "analz_idem";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   518
Addsimps [analz_idem];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   519
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   520
Goal "[| X\\<in> analz G;  G \\<subseteq> analz H |] ==> X\\<in> analz H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   521
by (dtac analz_mono 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   522
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   523
qed "analz_trans";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   524
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   525
(*Cut; Lemma 2 of Lowe*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   526
Goal "[| Y\\<in> analz (insert X H);  X\\<in> analz H |] ==> Y\\<in> analz H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   527
by (etac analz_trans 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   528
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   529
qed "analz_cut";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   530
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   531
(*Cut can be proved easily by induction on
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   532
   "Y: analz (insert X H) ==> X: analz H --> Y: analz H"
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   533
*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   534
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   535
(*This rewrite rule helps in the simplification of messages that involve
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   536
  the forwarding of unknown components (X).  Without it, removing occurrences
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   537
  of X can be very complicated. *)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   538
Goal "X\\<in> analz H ==> analz (insert X H) = analz H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   539
by (blast_tac (claset() addIs [analz_cut, analz_insertI]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   540
qed "analz_insert_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   541
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   542
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   543
(** A congruence rule for "analz" **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   544
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   545
Goal "[| analz G \\<subseteq> analz G'; analz H \\<subseteq> analz H' \
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   546
\              |] ==> analz (G Un H) \\<subseteq> analz (G' Un H')";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   547
by (Clarify_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   548
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   549
by (ALLGOALS (best_tac (claset() addIs [analz_mono RS subsetD])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   550
qed "analz_subset_cong";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   551
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   552
Goal "[| analz G = analz G'; analz H = analz H' \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   553
\              |] ==> analz (G Un H) = analz (G' Un H')";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   554
by (REPEAT_FIRST (ares_tac [equalityI, analz_subset_cong]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   555
          ORELSE' etac equalityE));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   556
qed "analz_cong";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   557
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   558
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   559
Goal "analz H = analz H' ==> analz(insert X H) = analz(insert X H')";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   560
by (asm_simp_tac (simpset() addsimps [insert_def] delsimps [singleton_conv]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   561
                            setloop (rtac analz_cong)) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   562
qed "analz_insert_cong";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   563
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   564
(*If there are no pairs or encryptions then analz does nothing*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   565
Goal "[| \\<forall>X Y. {|X,Y|} \\<notin> H;  \\<forall>X K. Crypt K X \\<notin> H |] ==> analz H = H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   566
by Safe_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   567
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   568
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   569
qed "analz_trivial";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   570
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   571
(*These two are obsolete (with a single Spy) but cost little to prove...*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   572
Goal "X\\<in> analz (\\<Union>i\\<in>A. analz (H i)) ==> X\\<in> analz (\\<Union>i\\<in>A. H i)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   573
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   574
by (ALLGOALS (blast_tac (claset() addIs [impOfSubs analz_mono])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   575
val lemma = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   576
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   577
Goal "analz (\\<Union>i\\<in>A. analz (H i)) = analz (\\<Union>i\\<in>A. H i)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   578
by (blast_tac (claset() addIs [lemma, impOfSubs analz_mono]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   579
qed "analz_UN_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   580
Addsimps [analz_UN_analz];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   581
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   582
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   583
(**** Inductive relation "synth" ****)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   584
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   585
Goal "H \\<subseteq> synth(H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   586
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   587
qed "synth_increasing";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   588
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   589
(** Unions **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   590
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   591
(*Converse fails: we can synth more from the union than from the 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   592
  separate parts, building a compound message using elements of each.*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   593
Goal "synth(G) Un synth(H) \\<subseteq> synth(G Un H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   594
by (REPEAT (ares_tac [Un_least, synth_mono, Un_upper1, Un_upper2] 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   595
qed "synth_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   596
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   597
Goal "insert X (synth H) \\<subseteq> synth(insert X H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   598
by (blast_tac (claset() addIs [impOfSubs synth_mono]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   599
qed "synth_insert";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   600
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   601
(** Idempotence and transitivity **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   602
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   603
Goal "X\\<in> synth (synth H) ==> X\\<in> synth H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   604
by (etac synth.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   605
by (ALLGOALS Blast_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   606
qed "synth_synthD";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   607
AddSDs [synth_synthD];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   608
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   609
Goal "synth (synth H) = synth H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   610
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   611
qed "synth_idem";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   612
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   613
Goal "[| X\\<in> synth G;  G \\<subseteq> synth H |] ==> X\\<in> synth H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   614
by (dtac synth_mono 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   615
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   616
qed "synth_trans";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   617
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   618
(*Cut; Lemma 2 of Lowe*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   619
Goal "[| Y\\<in> synth (insert X H);  X\\<in> synth H |] ==> Y\\<in> synth H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   620
by (etac synth_trans 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   621
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   622
qed "synth_cut";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   623
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   624
Goal "Agent A \\<in> synth H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   625
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   626
qed "Agent_synth";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   627
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   628
Goal "Number n \\<in> synth H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   629
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   630
qed "Number_synth";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   631
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   632
Goal "(Nonce N \\<in> synth H) = (Nonce N \\<in> H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   633
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   634
qed "Nonce_synth_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   635
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   636
Goal "(Key K \\<in> synth H) = (Key K \\<in> H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   637
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   638
qed "Key_synth_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   639
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   640
Goal "Key K \\<notin> H ==> (Crypt K X \\<in> synth H) = (Crypt K X \\<in> H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   641
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   642
qed "Crypt_synth_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   643
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   644
Addsimps [Agent_synth, Number_synth, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   645
	  Nonce_synth_eq, Key_synth_eq, Crypt_synth_eq];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   646
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   647
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   648
Goalw [keysFor_def]
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   649
    "keysFor (synth H) = keysFor H Un invKey`{K. Key K \\<in> H}";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   650
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   651
qed "keysFor_synth";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   652
Addsimps [keysFor_synth];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   653
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   654
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   655
(*** Combinations of parts, analz and synth ***)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   656
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   657
Goal "parts (synth H) = parts H Un synth H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   658
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   659
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   660
by (etac parts.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   661
by (ALLGOALS
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   662
    (blast_tac (claset() addIs [synth_increasing RS parts_mono RS subsetD,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   663
                                parts.Fst, parts.Snd, parts.Body])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   664
qed "parts_synth";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   665
Addsimps [parts_synth];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   666
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   667
Goal "analz (analz G Un H) = analz (G Un H)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   668
by (REPEAT_FIRST (resolve_tac [equalityI, analz_subset_cong]));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   669
by (ALLGOALS Simp_tac);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   670
qed "analz_analz_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   671
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   672
Goal "analz (synth G Un H) = analz (G Un H) Un synth G";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   673
by (rtac equalityI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   674
by (rtac subsetI 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   675
by (etac analz.induct 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   676
by (blast_tac (claset() addIs [impOfSubs analz_mono]) 5);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   677
by (ALLGOALS 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   678
    (blast_tac (claset() addIs [analz.Fst, analz.Snd, analz.Decrypt])));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   679
qed "analz_synth_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   680
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   681
Goal "analz (synth H) = analz H Un synth H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   682
by (cut_inst_tac [("H","{}")] analz_synth_Un 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   683
by (Full_simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   684
qed "analz_synth";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   685
Addsimps [analz_analz_Un, analz_synth_Un, analz_synth];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   686
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   687
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   688
(** For reasoning about the Fake rule in traces **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   689
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   690
Goal "X\\<in> G ==> parts(insert X H) \\<subseteq> parts G Un parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   691
by (rtac ([parts_mono, parts_Un_subset2] MRS subset_trans) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   692
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   693
qed "parts_insert_subset_Un";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   694
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   695
(*More specifically for Fake.  Very occasionally we could do with a version
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   696
  of the form  parts{X} \\<subseteq> synth (analz H) Un parts H *)
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   697
Goal "X\\<in> synth (analz H) ==> \
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   698
\     parts (insert X H) \\<subseteq> synth (analz H) Un parts H";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   699
by (dtac parts_insert_subset_Un 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   700
by (Full_simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   701
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   702
qed "Fake_parts_insert";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   703
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   704
(*H is sometimes (Key ` KK Un spies evs), so can't put G=H*)
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   705
Goal "X\\<in> synth (analz G) ==> \
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   706
\     analz (insert X H) \\<subseteq> synth (analz G) Un analz (G Un H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   707
by (rtac subsetI 1);
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   708
by (subgoal_tac "x \\<in> analz (synth (analz G) Un H)" 1);
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   709
by (blast_tac (claset() addIs [impOfSubs analz_mono,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   710
			       impOfSubs (analz_mono RS synth_mono)]) 2);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   711
by (Full_simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   712
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   713
qed "Fake_analz_insert";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   714
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   715
Goal "(X\\<in> analz H & X\\<in> parts H) = (X\\<in> analz H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   716
by (blast_tac (claset() addIs [impOfSubs analz_subset_parts]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   717
val analz_conj_parts = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   718
11251
a6816d47f41d converted many HOL/Auth theories to Isar scripts
paulson
parents: 11230
diff changeset
   719
Goal "(X\\<in> analz H | X\\<in> parts H) = (X\\<in> parts H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   720
by (blast_tac (claset() addIs [impOfSubs analz_subset_parts]) 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   721
val analz_disj_parts = result();
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   722
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   723
AddIffs [analz_conj_parts, analz_disj_parts];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   724
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   725
(*Without this equation, other rules for synth and analz would yield
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   726
  redundant cases*)
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   727
Goal "({|X,Y|} \\<in> synth (analz H)) = \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   728
\     (X \\<in> synth (analz H) & Y \\<in> synth (analz H))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   729
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   730
qed "MPair_synth_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   731
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   732
AddIffs [MPair_synth_analz];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   733
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   734
Goal "[| Key K \\<in> analz H;  Key (invKey K) \\<in> analz H |] \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   735
\      ==> (Crypt K X \\<in> synth (analz H)) = (X \\<in> synth (analz H))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   736
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   737
qed "Crypt_synth_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   738
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   739
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   740
Goal "X \\<notin> synth (analz H) \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   741
\     ==> (Hash{|X,Y|} \\<in> synth (analz H)) = (Hash{|X,Y|} \\<in> analz H)";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   742
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   743
qed "Hash_synth_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   744
Addsimps [Hash_synth_analz];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   745
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   746
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   747
(**** HPair: a combination of Hash and MPair ****)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   748
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   749
(*** Freeness ***)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   750
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   751
Goalw [HPair_def] "Agent A ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   752
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   753
qed "Agent_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   754
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   755
Goalw [HPair_def] "Nonce N ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   756
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   757
qed "Nonce_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   758
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   759
Goalw [HPair_def] "Number N ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   760
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   761
qed "Number_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   762
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   763
Goalw [HPair_def] "Key K ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   764
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   765
qed "Key_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   766
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   767
Goalw [HPair_def] "Hash Z ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   768
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   769
qed "Hash_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   770
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   771
Goalw [HPair_def] "Crypt K X' ~= Hash[X] Y";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   772
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   773
qed "Crypt_neq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   774
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   775
val HPair_neqs = [Agent_neq_HPair, Nonce_neq_HPair, Number_neq_HPair, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   776
                  Key_neq_HPair, Hash_neq_HPair, Crypt_neq_HPair];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   777
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   778
AddIffs HPair_neqs;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   779
AddIffs (HPair_neqs RL [not_sym]);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   780
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   781
Goalw [HPair_def] "(Hash[X'] Y' = Hash[X] Y) = (X' = X & Y'=Y)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   782
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   783
qed "HPair_eq";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   784
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   785
Goalw [HPair_def] "({|X',Y'|} = Hash[X] Y) = (X' = Hash{|X,Y|} & Y'=Y)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   786
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   787
qed "MPair_eq_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   788
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   789
Goalw [HPair_def] "(Hash[X] Y = {|X',Y'|}) = (X' = Hash{|X,Y|} & Y'=Y)";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   790
by Auto_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   791
qed "HPair_eq_MPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   792
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   793
AddIffs [HPair_eq, MPair_eq_HPair, HPair_eq_MPair];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   794
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   795
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   796
(*** Specialized laws, proved in terms of those for Hash and MPair ***)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   797
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   798
Goalw [HPair_def] "keysFor (insert (Hash[X] Y) H) = keysFor H";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   799
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   800
qed "keysFor_insert_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   801
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   802
Goalw [HPair_def]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   803
    "parts (insert (Hash[X] Y) H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   804
\    insert (Hash[X] Y) (insert (Hash{|X,Y|}) (parts (insert Y H)))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   805
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   806
qed "parts_insert_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   807
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   808
Goalw [HPair_def]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   809
    "analz (insert (Hash[X] Y) H) = \
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   810
\    insert (Hash[X] Y) (insert (Hash{|X,Y|}) (analz (insert Y H)))";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   811
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   812
qed "analz_insert_HPair";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   813
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   814
Goalw [HPair_def] "X \\<notin> synth (analz H) \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   815
\   ==> (Hash[X] Y \\<in> synth (analz H)) = \
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   816
\       (Hash {|X, Y|} \\<in> analz H & Y \\<in> synth (analz H))";
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   817
by (Simp_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   818
by (Blast_tac 1);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   819
qed "HPair_synth_analz";
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   820
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   821
Addsimps [keysFor_insert_HPair, parts_insert_HPair, analz_insert_HPair, 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   822
          HPair_synth_analz, HPair_synth_analz];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   823
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   824
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   825
(*We do NOT want Crypt... messages broken up in protocols!!*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   826
Delrules [make_elim parts.Body];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   827
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   828
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   829
(** Rewrites to push in Key and Crypt messages, so that other messages can
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   830
    be pulled out using the analz_insert rules **)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   831
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   832
fun insComm x y = inst "x" x (inst "y" y insert_commute);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   833
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   834
val pushKeys = map (insComm "Key ?K") 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   835
                   ["Agent ?C", "Nonce ?N", "Number ?N", 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   836
		    "Hash ?X", "MPair ?X ?Y", "Crypt ?X ?K'"];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   837
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   838
val pushCrypts = map (insComm "Crypt ?X ?K") 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   839
                     ["Agent ?C", "Nonce ?N", "Number ?N", 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   840
		      "Hash ?X'", "MPair ?X' ?Y"];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   841
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   842
(*Cannot be added with Addsimps -- we don't always want to re-order messages*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   843
bind_thms ("pushes", pushKeys@pushCrypts);
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   844
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   845
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   846
(*** Tactics useful for many protocol proofs ***)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   847
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   848
(*Prove base case (subgoal i) and simplify others.  A typical base case
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   849
  concerns  Crypt K X \\<notin> Key`shrK`bad  and cannot be proved by rewriting
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   850
  alone.*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   851
fun prove_simple_subgoals_tac i = 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   852
    force_tac (claset(), simpset() addsimps [image_eq_UN]) i THEN
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   853
    ALLGOALS Asm_simp_tac;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   854
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   855
fun Fake_parts_insert_tac i = 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   856
    blast_tac (claset() addIs [parts_insertI]
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   857
			addDs [impOfSubs analz_subset_parts,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   858
			       impOfSubs Fake_parts_insert]) i;
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   859
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   860
(*Apply rules to break down assumptions of the form
11217
54a86efcb0ba some X-symbols
paulson
parents: 11189
diff changeset
   861
  Y \\<in> parts(insert X H)  and  Y \\<in> analz(insert X H)
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   862
*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   863
val Fake_insert_tac = 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   864
    dresolve_tac [impOfSubs Fake_analz_insert,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   865
                  impOfSubs Fake_parts_insert] THEN'
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   866
    eresolve_tac [asm_rl, synth.Inj];
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   867
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   868
fun Fake_insert_simp_tac ss i = 
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   869
    REPEAT (Fake_insert_tac i) THEN asm_full_simp_tac ss i;
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   870
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   871
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   872
(*Analysis of Fake cases.  Also works for messages that forward unknown parts,
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   873
  but this application is no longer necessary if analz_insert_eq is used.
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   874
  Abstraction over i is ESSENTIAL: it delays the dereferencing of claset
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   875
  DEPENDS UPON "X" REFERRING TO THE FRADULENT MESSAGE *)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   876
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   877
fun atomic_spy_analz_tac (cs,ss) = SELECT_GOAL
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   878
    (Fake_insert_simp_tac ss 1
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   879
     THEN
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   880
     IF_UNSOLVED (Blast.depth_tac
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   881
		  (cs addIs [analz_insertI,
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   882
				   impOfSubs analz_subset_parts]) 4 1));
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   883
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   884
(*The explicit claset and simpset arguments help it work with Isar*)
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   885
fun gen_spy_analz_tac (cs,ss) i =
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   886
  DETERM
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   887
   (SELECT_GOAL
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   888
     (EVERY 
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   889
      [  (*push in occurrences of X...*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   890
       (REPEAT o CHANGED)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   891
           (res_inst_tac [("x1","X")] (insert_commute RS ssubst) 1),
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   892
       (*...allowing further simplifications*)
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   893
       simp_tac ss 1,
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   894
       REPEAT (FIRSTGOAL (resolve_tac [allI,impI,notI,conjI,iffI])),
11270
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   895
       DEPTH_SOLVE (atomic_spy_analz_tac (cs,ss) 1)]) i);
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   896
a315a3862bb4 better treatment of methods: uses Method.ctxt_args to refer to current
paulson
parents: 11264
diff changeset
   897
fun spy_analz_tac i = gen_spy_analz_tac (claset(), simpset()) i;
11189
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   898
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   899
(*By default only o_apply is built-in.  But in the presence of eta-expansion
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   900
  this means that some terms displayed as (f o g) will be rewritten, and others
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   901
  will not!*)
1ea763a5d186 conversion of Message.thy to Isar format
paulson
parents:
diff changeset
   902
Addsimps [o_def];