src/HOL/Auth/Guard/GuardK.thy
author wenzelm
Mon, 28 Dec 2015 23:13:33 +0100
changeset 61956 38b73f7940af
parent 61830 4f5ab843cf5b
child 62390 842917225d56
permissions -rw-r--r--
more symbols;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41775
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     1
(*  Title:      HOL/Auth/Guard/GuardK.thy
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     2
    Author:     Frederic Blanqui, University of Cambridge Computer Laboratory
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     3
    Copyright   2002  University of Cambridge
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     4
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     5
Very similar to Guard except:
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     6
- Guard is replaced by GuardK, guard by guardK, Nonce by Key
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     7
- some scripts are slightly modified (+ keyset_in, kparts_parts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     8
- the hypothesis Key n ~:G (keyset G) is added
41775
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     9
*)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    10
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    11
section\<open>protocol-independent confidentiality theorem on keys\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    12
27108
e447b3107696 whitespace tuning
haftmann
parents: 25134
diff changeset
    13
theory GuardK
e447b3107696 whitespace tuning
haftmann
parents: 25134
diff changeset
    14
imports Analz Extensions
e447b3107696 whitespace tuning
haftmann
parents: 25134
diff changeset
    15
begin
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    16
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    17
(******************************************************************************
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    18
messages where all the occurrences of Key n are
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    19
in a sub-message of the form Crypt (invKey K) X with K:Ks
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    20
******************************************************************************)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    21
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    22
inductive_set
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    23
  guardK :: "nat => key set => msg set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    24
  for n :: nat and Ks :: "key set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    25
where
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    26
  No_Key [intro]: "Key n ~:parts {X} ==> X:guardK n Ks"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    27
| Guard_Key [intro]: "invKey K:Ks ==> Crypt K X:guardK n Ks"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    28
| Crypt [intro]: "X:guardK n Ks ==> Crypt K X:guardK n Ks"
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
    29
| Pair [intro]: "[| X:guardK n Ks; Y:guardK n Ks |] ==> \<lbrace>X,Y\<rbrace>:guardK n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    30
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    31
subsection\<open>basic facts about @{term guardK}\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    32
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    33
lemma Nonce_is_guardK [iff]: "Nonce p:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    34
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    35
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    36
lemma Agent_is_guardK [iff]: "Agent A:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    37
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    38
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    39
lemma Number_is_guardK [iff]: "Number r:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    40
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    41
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    42
lemma Key_notin_guardK: "X:guardK n Ks ==> X ~= Key n"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    43
by (erule guardK.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    44
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    45
lemma Key_notin_guardK_iff [iff]: "Key n ~:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    46
by (auto dest: Key_notin_guardK)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    47
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    48
lemma guardK_has_Crypt [rule_format]: "X:guardK n Ks ==> Key n:parts {X}
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    49
--> (EX K Y. Crypt K Y:kparts {X} & Key n:parts {Y})"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    50
by (erule guardK.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    51
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    52
lemma Key_notin_kparts_msg: "X:guardK n Ks ==> Key n ~:kparts {X}"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    53
by (erule guardK.induct, auto dest: kparts_parts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    54
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    55
lemma Key_in_kparts_imp_no_guardK: "Key n:kparts H
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    56
==> EX X. X:H & X ~:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    57
apply (drule in_kparts, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    58
apply (rule_tac x=X in exI, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    59
by (auto dest: Key_notin_kparts_msg)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    60
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    61
lemma guardK_kparts [rule_format]: "X:guardK n Ks ==>
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    62
Y:kparts {X} --> Y:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    63
by (erule guardK.induct, auto dest: kparts_parts parts_sub)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    64
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    65
lemma guardK_Crypt: "[| Crypt K Y:guardK n Ks; K ~:invKey`Ks |] ==> Y:guardK n Ks"
56681
e8d5d60d655e avoid non-standard simp default rule
haftmann
parents: 55417
diff changeset
    66
  by (ind_cases "Crypt K Y:guardK n Ks") (auto intro!: image_eqI)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    67
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
    68
lemma guardK_MPair [iff]: "(\<lbrace>X,Y\<rbrace>:guardK n Ks)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    69
= (X:guardK n Ks & Y:guardK n Ks)"
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
    70
by (auto, (ind_cases "\<lbrace>X,Y\<rbrace>:guardK n Ks", auto)+)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    71
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    72
lemma guardK_not_guardK [rule_format]: "X:guardK n Ks ==>
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    73
Crypt K Y:kparts {X} --> Key n:kparts {Y} --> Y ~:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    74
by (erule guardK.induct, auto dest: guardK_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    75
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    76
lemma guardK_extand: "[| X:guardK n Ks; Ks <= Ks';
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    77
[| K:Ks'; K ~:Ks |] ==> Key K ~:parts {X} |] ==> X:guardK n Ks'"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    78
by (erule guardK.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    79
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    80
subsection\<open>guarded sets\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    81
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 27108
diff changeset
    82
definition GuardK :: "nat => key set => msg set => bool" where
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    83
"GuardK n Ks H == ALL X. X:H --> X:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    84
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    85
subsection\<open>basic facts about @{term GuardK}\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    86
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    87
lemma GuardK_empty [iff]: "GuardK n Ks {}"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    88
by (simp add: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    89
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    90
lemma Key_notin_kparts [simplified]: "GuardK n Ks H ==> Key n ~:kparts H"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    91
by (auto simp: GuardK_def dest: in_kparts Key_notin_kparts_msg)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    92
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    93
lemma GuardK_must_decrypt: "[| GuardK n Ks H; Key n:analz H |] ==>
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    94
EX K Y. Crypt K Y:kparts H & Key (invKey K):kparts H"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    95
apply (drule_tac P="%G. Key n:G" in analz_pparts_kparts_substD, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    96
by (drule must_decrypt, auto dest: Key_notin_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    97
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    98
lemma GuardK_kparts [intro]: "GuardK n Ks H ==> GuardK n Ks (kparts H)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    99
by (auto simp: GuardK_def dest: in_kparts guardK_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   100
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   101
lemma GuardK_mono: "[| GuardK n Ks H; G <= H |] ==> GuardK n Ks G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   102
by (auto simp: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   103
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   104
lemma GuardK_insert [iff]: "GuardK n Ks (insert X H)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   105
= (GuardK n Ks H & X:guardK n Ks)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   106
by (auto simp: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   107
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   108
lemma GuardK_Un [iff]: "GuardK n Ks (G Un H) = (GuardK n Ks G & GuardK n Ks H)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   109
by (auto simp: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   110
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   111
lemma GuardK_synth [intro]: "GuardK n Ks G ==> GuardK n Ks (synth G)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   112
by (auto simp: GuardK_def, erule synth.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   113
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   114
lemma GuardK_analz [intro]: "[| GuardK n Ks G; ALL K. K:Ks --> Key K ~:analz G |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   115
==> GuardK n Ks (analz G)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   116
apply (auto simp: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   117
apply (erule analz.induct, auto)
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
   118
by (ind_cases "Crypt K Xa:guardK n Ks" for K Xa, auto)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   119
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   120
lemma in_GuardK [dest]: "[| X:G; GuardK n Ks G |] ==> X:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   121
by (auto simp: GuardK_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   122
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   123
lemma in_synth_GuardK: "[| X:synth G; GuardK n Ks G |] ==> X:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   124
by (drule GuardK_synth, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   125
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   126
lemma in_analz_GuardK: "[| X:analz G; GuardK n Ks G;
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   127
ALL K. K:Ks --> Key K ~:analz G |] ==> X:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   128
by (drule GuardK_analz, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   129
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   130
lemma GuardK_keyset [simp]: "[| keyset G; Key n ~:G |] ==> GuardK n Ks G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   131
by (simp only: GuardK_def, clarify, drule keyset_in, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   132
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   133
lemma GuardK_Un_keyset: "[| GuardK n Ks G; keyset H; Key n ~:H |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   134
==> GuardK n Ks (G Un H)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   135
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   136
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   137
lemma in_GuardK_kparts: "[| X:G; GuardK n Ks G; Y:kparts {X} |] ==> Y:guardK n Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   138
by blast
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   139
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   140
lemma in_GuardK_kparts_neq: "[| X:G; GuardK n Ks G; Key n':kparts {X} |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   141
==> n ~= n'"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   142
by (blast dest: in_GuardK_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   143
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   144
lemma in_GuardK_kparts_Crypt: "[| X:G; GuardK n Ks G; is_MPair X;
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   145
Crypt K Y:kparts {X}; Key n:kparts {Y} |] ==> invKey K:Ks"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   146
apply (drule in_GuardK, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   147
apply (frule guardK_not_guardK, simp+)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   148
apply (drule guardK_kparts, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   149
by (ind_cases "Crypt K Y:guardK n Ks", auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   150
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   151
lemma GuardK_extand: "[| GuardK n Ks G; Ks <= Ks';
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   152
[| K:Ks'; K ~:Ks |] ==> Key K ~:parts G |] ==> GuardK n Ks' G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   153
by (auto simp: GuardK_def dest: guardK_extand parts_sub)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   154
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   155
subsection\<open>set obtained by decrypting a message\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   156
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 19233
diff changeset
   157
abbreviation (input)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
   158
  decrypt :: "msg set => key => msg => msg set" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 19233
diff changeset
   159
  "decrypt H K Y == insert Y (H - {Crypt K Y})"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   160
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   161
lemma analz_decrypt: "[| Crypt K Y:H; Key (invKey K):H; Key n:analz H |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   162
==> Key n:analz (decrypt H K Y)"
14307
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   163
apply (drule_tac P="%H. Key n:analz H" in ssubst [OF insert_Diff])
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   164
apply assumption 
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   165
apply (simp only: analz_Crypt_if, simp)
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   166
done
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   167
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   168
lemma parts_decrypt: "[| Crypt K Y:H; X:parts (decrypt H K Y) |] ==> X:parts H"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   169
by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   170
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   171
subsection\<open>number of Crypt's in a message\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   172
35418
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   173
fun crypt_nb :: "msg => nat" where
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   174
"crypt_nb (Crypt K X) = Suc (crypt_nb X)" |
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
   175
"crypt_nb \<lbrace>X,Y\<rbrace> = crypt_nb X + crypt_nb Y" |
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   176
"crypt_nb X = 0" (* otherwise *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   177
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   178
subsection\<open>basic facts about @{term crypt_nb}\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   179
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 23746
diff changeset
   180
lemma non_empty_crypt_msg: "Crypt K Y:parts {X} ==> crypt_nb X \<noteq> 0"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   181
by (induct X, simp_all, safe, simp_all)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   182
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   183
subsection\<open>number of Crypt's in a message list\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   184
35418
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   185
primrec cnb :: "msg list => nat" where
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   186
"cnb [] = 0" |
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   187
"cnb (X#l) = crypt_nb X + cnb l"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   188
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   189
subsection\<open>basic facts about @{term cnb}\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   190
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   191
lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   192
by (induct l, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   193
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   194
lemma mem_cnb_minus: "x \<in> set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   195
by (induct l, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   196
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   197
lemmas mem_cnb_minus_substI = mem_cnb_minus [THEN ssubst]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   198
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   199
lemma cnb_minus [simp]: "x \<in> set l ==> cnb (remove l x) = cnb l - crypt_nb x"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   200
apply (induct l, auto)
45600
1bbbac9a0cb0 'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents: 41775
diff changeset
   201
by (erule_tac l=l and x=x in mem_cnb_minus_substI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   202
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   203
lemma parts_cnb: "Z:parts (set l) ==>
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   204
cnb l = (cnb l - crypt_nb Z) + crypt_nb Z"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   205
by (erule parts.induct, auto simp: in_set_conv_decomp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   206
25134
3d4953e88449 Eliminated most of the neq0_conv occurrences. As a result, many
nipkow
parents: 23746
diff changeset
   207
lemma non_empty_crypt: "Crypt K Y:parts (set l) ==> cnb l \<noteq> 0"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   208
by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   209
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   210
subsection\<open>list of kparts\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   211
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   212
lemma kparts_msg_set: "EX l. kparts {X} = set l & cnb l = crypt_nb X"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   213
apply (induct X, simp_all)
58250
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   214
apply (rename_tac agent, rule_tac x="[Agent agent]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   215
apply (rename_tac nat, rule_tac x="[Number nat]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   216
apply (rename_tac nat, rule_tac x="[Nonce nat]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   217
apply (rename_tac nat, rule_tac x="[Key nat]" in exI, simp)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 14307
diff changeset
   218
apply (rule_tac x="[Hash X]" in exI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   219
apply (clarify, rule_tac x="l@la" in exI, simp)
58250
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   220
by (clarify, rename_tac nat X y, rule_tac x="[Crypt nat X]" in exI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   221
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   222
lemma kparts_set: "EX l'. kparts (set l) = set l' & cnb l' = cnb l"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   223
apply (induct l)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   224
apply (rule_tac x="[]" in exI, simp, clarsimp)
55417
01fbfb60c33e adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents: 45600
diff changeset
   225
apply (rename_tac a b l')
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 14307
diff changeset
   226
apply (subgoal_tac "EX l''.  kparts {a} = set l'' & cnb l'' = crypt_nb a", clarify)
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 14307
diff changeset
   227
apply (rule_tac x="l''@l'" in exI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   228
apply (rule kparts_insert_substI, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   229
by (rule kparts_msg_set)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   230
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   231
subsection\<open>list corresponding to "decrypt"\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   232
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 27108
diff changeset
   233
definition decrypt' :: "msg list => key => msg => msg list" where
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 17087
diff changeset
   234
"decrypt' l K Y == Y # remove l (Crypt K Y)"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   235
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   236
declare decrypt'_def [simp]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   237
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   238
subsection\<open>basic facts about @{term decrypt'}\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   239
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   240
lemma decrypt_minus: "decrypt (set l) K Y <= set (decrypt' l K Y)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   241
by (induct l, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   242
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   243
text\<open>if the analysis of a finite guarded set gives n then it must also give
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   244
one of the keys of Ks\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   245
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   246
lemma GuardK_invKey_by_list [rule_format]: "ALL l. cnb l = p
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   247
--> GuardK n Ks (set l) --> Key n:analz (set l)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   248
--> (EX K. K:Ks & Key K:analz (set l))"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   249
apply (induct p)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   250
(* case p=0 *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   251
apply (clarify, drule GuardK_must_decrypt, simp, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   252
apply (drule kparts_parts, drule non_empty_crypt, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   253
(* case p>0 *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   254
apply (clarify, frule GuardK_must_decrypt, simp, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   255
apply (drule_tac P="%G. Key n:G" in analz_pparts_kparts_substD, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   256
apply (frule analz_decrypt, simp_all)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   257
apply (subgoal_tac "EX l'. kparts (set l) = set l' & cnb l' = cnb l", clarsimp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   258
apply (drule_tac G="insert Y (set l' - {Crypt K Y})"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   259
and H="set (decrypt' l' K Y)" in analz_sub, rule decrypt_minus)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   260
apply (rule_tac analz_pparts_kparts_substI, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   261
apply (case_tac "K:invKey`Ks")
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   262
(* K:invKey`Ks *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   263
apply (clarsimp, blast)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   264
(* K ~:invKey`Ks *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   265
apply (subgoal_tac "GuardK n Ks (set (decrypt' l' K Y))")
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   266
apply (drule_tac x="decrypt' l' K Y" in spec, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   267
apply (subgoal_tac "Crypt K Y:parts (set l)")
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   268
apply (drule parts_cnb, rotate_tac -1, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   269
apply (clarify, drule_tac X="Key Ka" and H="insert Y (set l')" in analz_sub)
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 17087
diff changeset
   270
apply (rule insert_mono, rule set_remove)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   271
apply (simp add: analz_insertD, blast)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   272
(* Crypt K Y:parts (set l) *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   273
apply (blast dest: kparts_parts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   274
(* GuardK n Ks (set (decrypt' l' K Y)) *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   275
apply (rule_tac H="insert Y (set l')" in GuardK_mono)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   276
apply (subgoal_tac "GuardK n Ks (set l')", simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   277
apply (rule_tac K=K in guardK_Crypt, simp add: GuardK_def, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   278
apply (drule_tac t="set l'" in sym, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   279
apply (rule GuardK_kparts, simp, simp)
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 17087
diff changeset
   280
apply (rule_tac B="set l'" in subset_trans, rule set_remove, blast)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   281
by (rule kparts_set)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   282
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   283
lemma GuardK_invKey_finite: "[| Key n:analz G; GuardK n Ks G; finite G |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   284
==> EX K. K:Ks & Key K:analz G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   285
apply (drule finite_list, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   286
by (rule GuardK_invKey_by_list, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   287
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   288
lemma GuardK_invKey: "[| Key n:analz G; GuardK n Ks G |]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   289
==> EX K. K:Ks & Key K:analz G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   290
by (auto dest: analz_needs_only_finite GuardK_invKey_finite)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   291
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   292
text\<open>if the analyse of a finite guarded set and a (possibly infinite) set of
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   293
keys gives n then it must also gives Ks\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   294
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   295
lemma GuardK_invKey_keyset: "[| Key n:analz (G Un H); GuardK n Ks G; finite G;
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   296
keyset H; Key n ~:H |] ==> EX K. K:Ks & Key K:analz (G Un H)"
45600
1bbbac9a0cb0 'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents: 41775
diff changeset
   297
apply (frule_tac P="%G. Key n:G" and G=G in analz_keyset_substD, simp_all)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   298
apply (drule_tac G="G Un (H Int keysfor G)" in GuardK_invKey_finite)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   299
apply (auto simp: GuardK_def intro: analz_sub)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   300
by (drule keyset_in, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   301
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   302
end