src/HOL/Auth/Guard/Guard.thy
author wenzelm
Mon, 13 Apr 2020 22:08:14 +0200
changeset 71751 abf3e80bd815
parent 69597 ff784d5a5bfb
child 76287 cdc14f94c754
permissions -rw-r--r--
tuned NEWS;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41775
6214816d79d3 standardized headers;
wenzelm
parents: 37596
diff changeset
     1
(*  Title:      HOL/Auth/Guard/Guard.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
*)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     5
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
     6
section\<open>Protocol-Independent Confidentiality Theorem on Nonces\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 15236
diff changeset
     8
theory Guard imports Analz Extensions begin
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
     9
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    10
(******************************************************************************
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    11
messages where all the occurrences of Nonce n are
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    12
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
    13
******************************************************************************)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    14
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    15
inductive_set
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    16
  guard :: "nat \<Rightarrow> key set \<Rightarrow> msg set"
23746
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    17
  for n :: nat and Ks :: "key set"
a455e69c31cc Adapted to new inductive definition package.
berghofe
parents: 21404
diff changeset
    18
where
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    19
  No_Nonce [intro]: "Nonce n \<notin> parts {X} \<Longrightarrow> X \<in> guard n Ks"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    20
| Guard_Nonce [intro]: "invKey K \<in> Ks \<Longrightarrow> Crypt K X \<in> guard n Ks"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    21
| Crypt [intro]: "X \<in> guard n Ks \<Longrightarrow> Crypt K X \<in> guard n Ks"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    22
| Pair [intro]: "[| X \<in> guard n Ks; Y \<in> guard n Ks |] ==> \<lbrace>X,Y\<rbrace> \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    23
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    24
subsection\<open>basic facts about \<^term>\<open>guard\<close>\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    25
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    26
lemma Key_is_guard [iff]: "Key K \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    27
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    28
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    29
lemma Agent_is_guard [iff]: "Agent A \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    30
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    31
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    32
lemma Number_is_guard [iff]: "Number r \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    33
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    34
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    35
lemma Nonce_notin_guard: "X \<in> guard n Ks \<Longrightarrow> X \<noteq> Nonce n"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    36
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    37
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    38
lemma Nonce_notin_guard_iff [iff]: "Nonce n \<notin> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    39
by (auto dest: Nonce_notin_guard)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    40
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    41
lemma guard_has_Crypt [rule_format]: "X \<in> guard n Ks ==> Nonce n \<in> parts {X}
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    42
\<longrightarrow> (\<exists>K Y. Crypt K Y \<in> kparts {X} \<and> Nonce n \<in> parts {Y})"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    43
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    44
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    45
lemma Nonce_notin_kparts_msg: "X \<in> guard n Ks \<Longrightarrow> Nonce n \<notin> kparts {X}"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    46
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    47
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    48
lemma Nonce_in_kparts_imp_no_guard: "Nonce n \<in> kparts H
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    49
\<Longrightarrow> \<exists>X. X \<in> H \<and> X \<notin> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    50
apply (drule in_kparts, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    51
apply (rule_tac x=X in exI, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    52
by (auto dest: Nonce_notin_kparts_msg)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    53
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    54
lemma guard_kparts [rule_format]: "X \<in> guard n Ks \<Longrightarrow>
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    55
Y \<in> kparts {X} \<longrightarrow> Y \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    56
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    57
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    58
lemma guard_Crypt: "[| Crypt K Y \<in> guard n Ks; K \<notin> invKey`Ks |] ==> Y \<in> guard n Ks"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    59
  by (ind_cases "Crypt K Y \<in> guard n Ks") (auto intro!: image_eqI)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    60
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
    61
lemma guard_MPair [iff]: "(\<lbrace>X,Y\<rbrace> \<in> guard n Ks) = (X \<in> guard n Ks \<and> Y \<in> guard n Ks)"
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
    62
by (auto, (ind_cases "\<lbrace>X,Y\<rbrace> \<in> guard n Ks", auto)+)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    63
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    64
lemma guard_not_guard [rule_format]: "X \<in> guard n Ks \<Longrightarrow>
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    65
Crypt K Y \<in> kparts {X} \<longrightarrow> Nonce n \<in> kparts {Y} \<longrightarrow> Y \<notin> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    66
by (erule guard.induct, auto dest: guard_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    67
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    68
lemma guard_extand: "[| X \<in> guard n Ks; Ks \<subseteq> Ks' |] ==> X \<in> guard n Ks'"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    69
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    70
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
    71
subsection\<open>guarded sets\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    72
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    73
definition Guard :: "nat \<Rightarrow> key set \<Rightarrow> msg set \<Rightarrow> bool" where
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    74
"Guard n Ks H \<equiv> \<forall>X. X \<in> H \<longrightarrow> X \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    75
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    76
subsection\<open>basic facts about \<^term>\<open>Guard\<close>\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    77
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    78
lemma Guard_empty [iff]: "Guard n Ks {}"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    79
by (simp add: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    80
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    81
lemma notin_parts_Guard [intro]: "Nonce n \<notin> parts G \<Longrightarrow> Guard n Ks G"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    82
apply (unfold Guard_def, clarify)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    83
apply (subgoal_tac "Nonce n \<notin> parts {X}")
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    84
by (auto dest: parts_sub)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    85
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    86
lemma Nonce_notin_kparts [simplified]: "Guard n Ks H \<Longrightarrow> Nonce n \<notin> kparts H"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    87
by (auto simp: Guard_def dest: in_kparts Nonce_notin_kparts_msg)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    88
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    89
lemma Guard_must_decrypt: "[| Guard n Ks H; Nonce n \<in> analz H |] ==>
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    90
\<exists>K Y. Crypt K Y \<in> kparts H \<and> Key (invKey K) \<in> kparts H"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
    91
apply (drule_tac P="\<lambda>G. Nonce n \<in> G" in analz_pparts_kparts_substD, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    92
by (drule must_decrypt, auto dest: Nonce_notin_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    93
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    94
lemma Guard_kparts [intro]: "Guard n Ks H ==> Guard n Ks (kparts H)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    95
by (auto simp: Guard_def dest: in_kparts guard_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    96
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    97
lemma Guard_mono: "[| Guard n Ks H; G <= H |] ==> Guard n Ks G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    98
by (auto simp: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
    99
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   100
lemma Guard_insert [iff]: "Guard n Ks (insert X H)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   101
= (Guard n Ks H \<and> X \<in> guard n Ks)"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   102
by (auto simp: Guard_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 Guard_Un [iff]: "Guard n Ks (G Un H) = (Guard n Ks G & Guard n Ks H)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   105
by (auto simp: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   106
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   107
lemma Guard_synth [intro]: "Guard n Ks G ==> Guard n Ks (synth G)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   108
by (auto simp: Guard_def, erule synth.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   109
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   110
lemma Guard_analz [intro]: "[| Guard n Ks G; \<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |]
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   111
==> Guard n Ks (analz G)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   112
apply (auto simp: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   113
apply (erule analz.induct, auto)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   114
by (ind_cases "Crypt K Xa \<in> guard n Ks" for K Xa, auto)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   115
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   116
lemma in_Guard [dest]: "[| X \<in> G; Guard n Ks G |] ==> X \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   117
by (auto simp: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   118
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   119
lemma in_synth_Guard: "[| X \<in> synth G; Guard n Ks G |] ==> X \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   120
by (drule Guard_synth, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   121
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   122
lemma in_analz_Guard: "[| X \<in> analz G; Guard n Ks G;
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   123
\<forall>K. K \<in> Ks \<longrightarrow> Key K \<notin> analz G |] ==> X \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   124
by (drule Guard_analz, 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 Guard_keyset [simp]: "keyset G ==> Guard n Ks G"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   127
by (auto simp: Guard_def)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   128
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   129
lemma Guard_Un_keyset: "[| Guard n Ks G; keyset H |] ==> Guard n Ks (G \<union> H)"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   130
by auto
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   131
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   132
lemma in_Guard_kparts: "[| X \<in> G; Guard n Ks G; Y \<in> kparts {X} |] ==> Y \<in> guard n Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   133
by blast
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   134
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   135
lemma in_Guard_kparts_neq: "[| X \<in> G; Guard n Ks G; Nonce n' \<in> kparts {X} |]
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   136
==> n \<noteq> n'"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   137
by (blast dest: in_Guard_kparts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   138
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   139
lemma in_Guard_kparts_Crypt: "[| X \<in> G; Guard n Ks G; is_MPair X;
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   140
Crypt K Y \<in> kparts {X}; Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   141
apply (drule in_Guard, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   142
apply (frule guard_not_guard, simp+)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   143
apply (drule guard_kparts, simp)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   144
by (ind_cases "Crypt K Y \<in> guard n Ks", auto)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   145
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   146
lemma Guard_extand: "[| Guard n Ks G; Ks \<subseteq> Ks' |] ==> Guard n Ks' G"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   147
by (auto simp: Guard_def dest: guard_extand)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   148
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   149
lemma guard_invKey [rule_format]: "[| X \<in> guard n Ks; Nonce n \<in> kparts {Y} |] ==>
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   150
Crypt K Y \<in> kparts {X} \<longrightarrow> invKey K \<in> Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   151
by (erule guard.induct, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   152
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   153
lemma Crypt_guard_invKey [rule_format]: "[| Crypt K Y \<in> guard n Ks;
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   154
Nonce n \<in> kparts {Y} |] ==> invKey K \<in> Ks"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   155
by (auto dest: guard_invKey)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   156
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   157
subsection\<open>set obtained by decrypting a message\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   158
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 19233
diff changeset
   159
abbreviation (input)
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 20768
diff changeset
   160
  decrypt :: "msg set => key => msg => msg set" where
20768
1d478c2d621f replaced syntax/translations by abbreviation;
wenzelm
parents: 19233
diff changeset
   161
  "decrypt H K Y == insert Y (H - {Crypt K Y})"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   162
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   163
lemma analz_decrypt: "[| Crypt K Y \<in> H; Key (invKey K) \<in> H; Nonce n \<in> analz H |]
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   164
==> Nonce n \<in> analz (decrypt H K Y)"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   165
apply (drule_tac P="\<lambda>H. Nonce n \<in> analz H" in ssubst [OF insert_Diff])
14307
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   166
apply assumption
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   167
apply (simp only: analz_Crypt_if, simp)
1cbc24648cf7 tidying of HOL/Auth esp Guard lemmas
paulson
parents: 13508
diff changeset
   168
done
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   169
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   170
lemma parts_decrypt: "[| Crypt K Y \<in> H; X \<in> parts (decrypt H K Y) |] ==> X \<in> parts H"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   171
by (erule parts.induct, auto intro: parts.Fst parts.Snd parts.Body)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   172
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   173
subsection\<open>number of Crypt's in a message\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   174
35418
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   175
fun crypt_nb :: "msg => nat"
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   176
where
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   177
  "crypt_nb (Crypt K X) = Suc (crypt_nb X)"
61956
38b73f7940af more symbols;
wenzelm
parents: 61830
diff changeset
   178
| "crypt_nb \<lbrace>X,Y\<rbrace> = crypt_nb X + crypt_nb Y"
35418
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   179
| "crypt_nb X = 0" (* otherwise *)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   180
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   181
subsection\<open>basic facts about \<^term>\<open>crypt_nb\<close>\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   182
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   183
lemma non_empty_crypt_msg: "Crypt K Y \<in> parts {X} \<Longrightarrow> crypt_nb X \<noteq> 0"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   184
by (induct X, simp_all, safe, simp_all)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   185
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   186
subsection\<open>number of Crypt's in a message list\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   187
35418
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   188
primrec cnb :: "msg list => nat"
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   189
where
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   190
  "cnb [] = 0"
83b0f75810f0 killed recdefs in HOL-Auth
krauss
parents: 35416
diff changeset
   191
| "cnb (X#l) = crypt_nb X + cnb l"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   192
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   193
subsection\<open>basic facts about \<^term>\<open>cnb\<close>\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   194
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   195
lemma cnb_app [simp]: "cnb (l @ l') = cnb l + cnb l'"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   196
by (induct l, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   197
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   198
lemma mem_cnb_minus: "x \<in> set l ==> cnb l = crypt_nb x + (cnb l - crypt_nb x)"
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   199
  by (induct l) auto
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   200
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   201
lemmas mem_cnb_minus_substI = mem_cnb_minus [THEN ssubst]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   202
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   203
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
   204
apply (induct l, auto)
45600
1bbbac9a0cb0 'lemmas' / 'theorems' commands allow 'for' fixes and standardize the result before storing;
wenzelm
parents: 41775
diff changeset
   205
apply (erule_tac l=l and x=x in mem_cnb_minus_substI)
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   206
apply simp
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   207
done
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   208
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   209
lemma parts_cnb: "Z \<in> parts (set l) \<Longrightarrow>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   210
cnb l = (cnb l - crypt_nb Z) + crypt_nb Z"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   211
by (erule parts.induct, auto simp: in_set_conv_decomp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   212
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   213
lemma non_empty_crypt: "Crypt K Y \<in> parts (set l) \<Longrightarrow> cnb l \<noteq> 0"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   214
by (induct l, auto dest: non_empty_crypt_msg parts_insert_substD)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   215
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   216
subsection\<open>list of kparts\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   217
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   218
lemma kparts_msg_set: "\<exists>l. kparts {X} = set l \<and> cnb l = crypt_nb X"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   219
apply (induct X, simp_all)
58250
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   220
apply (rename_tac agent, rule_tac x="[Agent agent]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   221
apply (rename_tac nat, rule_tac x="[Number nat]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   222
apply (rename_tac nat, rule_tac x="[Nonce nat]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   223
apply (rename_tac nat, rule_tac x="[Key nat]" in exI, simp)
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   224
apply (rename_tac X, rule_tac x="[Hash X]" in exI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   225
apply (clarify, rule_tac x="l@la" in exI, simp)
58250
bf4188deabb2 rename_tac'd scripts
blanchet
parents: 56681
diff changeset
   226
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
   227
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   228
lemma kparts_set: "\<exists>l'. kparts (set l) = set l' \<and> cnb l' = cnb l"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   229
apply (induct l)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   230
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
   231
apply (rename_tac a b l')
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   232
apply (subgoal_tac "\<exists>l''.  kparts {a} = set l'' \<and> cnb l'' = crypt_nb a", clarify)
15236
f289e8ba2bb3 Proofs needed to be updated because induction now preserves name of
nipkow
parents: 14307
diff changeset
   233
apply (rule_tac x="l''@l'" in exI, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   234
apply (rule kparts_insert_substI, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   235
by (rule kparts_msg_set)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   236
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   237
subsection\<open>list corresponding to "decrypt"\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   238
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 25134
diff changeset
   239
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
   240
"decrypt' l K Y == Y # remove l (Crypt K Y)"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   241
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   242
declare decrypt'_def [simp]
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   243
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   244
subsection\<open>basic facts about \<^term>\<open>decrypt'\<close>\<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 decrypt_minus: "decrypt (set l) K Y <= set (decrypt' l K Y)"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   247
by (induct l, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   248
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   249
subsection\<open>if the analyse of a finite guarded set gives n then it must also gives
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   250
one of the keys of Ks\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   251
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   252
lemma Guard_invKey_by_list [rule_format]: "\<forall>l. cnb l = p
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   253
\<longrightarrow> Guard n Ks (set l) \<longrightarrow> Nonce n \<in> analz (set l)
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   254
\<longrightarrow> (\<exists>K. K \<in> Ks \<and> Key K \<in> analz (set l))"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   255
apply (induct p)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   256
(* case p=0 *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   257
apply (clarify, drule Guard_must_decrypt, simp, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   258
apply (drule kparts_parts, drule non_empty_crypt, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   259
(* case p>0 *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   260
apply (clarify, frule Guard_must_decrypt, simp, clarify)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   261
apply (drule_tac P="\<lambda>G. Nonce n \<in> G" in analz_pparts_kparts_substD, simp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   262
apply (frule analz_decrypt, simp_all)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   263
apply (subgoal_tac "\<exists>l'. kparts (set l) = set l' \<and> cnb l' = cnb l", clarsimp)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   264
apply (drule_tac G="insert Y (set l' - {Crypt K Y})"
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   265
and H="set (decrypt' l' K Y)" in analz_sub, rule decrypt_minus)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   266
apply (rule_tac analz_pparts_kparts_substI, simp)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   267
apply (case_tac "K \<in> invKey`Ks")
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   268
(* K:invKey`Ks *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   269
apply (clarsimp, blast)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   270
(* K ~:invKey`Ks *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   271
apply (subgoal_tac "Guard n Ks (set (decrypt' l' K Y))")
37596
248db70c9bcf dropped ancient infix mem
haftmann
parents: 35418
diff changeset
   272
apply (drule_tac x="decrypt' l' K Y" in spec, simp)
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   273
apply (subgoal_tac "Crypt K Y \<in> parts (set l)")
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   274
apply (drule parts_cnb, rotate_tac -1, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   275
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
   276
apply (rule insert_mono, rule set_remove)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   277
apply (simp add: analz_insertD, blast)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   278
(* Crypt K Y:parts (set l) *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   279
apply (blast dest: kparts_parts)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   280
(* Guard n Ks (set (decrypt' l' K Y)) *)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   281
apply (rule_tac H="insert Y (set l')" in Guard_mono)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   282
apply (subgoal_tac "Guard n Ks (set l')", simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   283
apply (rule_tac K=K in guard_Crypt, simp add: Guard_def, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   284
apply (drule_tac t="set l'" in sym, simp)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   285
apply (rule Guard_kparts, simp, simp)
19233
77ca20b0ed77 renamed HOL + - * etc. to HOL.plus HOL.minus HOL.times etc.
haftmann
parents: 17087
diff changeset
   286
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
   287
by (rule kparts_set)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   288
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   289
lemma Guard_invKey_finite: "[| Nonce n \<in> analz G; Guard n Ks G; finite G |]
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   290
==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   291
apply (drule finite_list, clarify)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   292
by (rule Guard_invKey_by_list, auto)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   293
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   294
lemma Guard_invKey: "[| Nonce n \<in> analz G; Guard n Ks G |]
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   295
==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz G"
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   296
by (auto dest: analz_needs_only_finite Guard_invKey_finite)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   297
61830
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   298
subsection\<open>if the analyse of a finite guarded set and a (possibly infinite) set of keys
4f5ab843cf5b isabelle update_cartouches -c -t;
wenzelm
parents: 58889
diff changeset
   299
gives n then it must also gives Ks\<close>
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   300
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   301
lemma Guard_invKey_keyset: "[| Nonce n \<in> analz (G \<union> H); Guard n Ks G; finite G;
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   302
keyset H |] ==> \<exists>K. K \<in> Ks \<and> Key K \<in> analz (G \<union> H)"
ce654b0e6d69 more symbols;
wenzelm
parents: 62390
diff changeset
   303
apply (frule_tac P="\<lambda>G. Nonce n \<in> G" and G=G in analz_keyset_substD, simp_all)
13508
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   304
apply (drule_tac G="G Un (H Int keysfor G)" in Guard_invKey_finite)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   305
by (auto simp: Guard_def intro: analz_sub)
890d736b93a5 Frederic Blanqui's new "guard" examples
paulson
parents:
diff changeset
   306
62390
842917225d56 more canonical names
nipkow
parents: 61956
diff changeset
   307
end