author | wenzelm |
Mon, 08 Jun 2020 21:55:14 +0200 | |
changeset 71926 | bee83c9d3306 |
parent 69597 | ff784d5a5bfb |
permissions | -rw-r--r-- |
33028 | 1 |
(* Title: HOL/SET_Protocol/Event_SET.thy |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
2 |
Author: Giampaolo Bella |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
3 |
Author: Fabio Massacci |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
4 |
Author: Lawrence C Paulson |
14199 | 5 |
*) |
6 |
||
63167 | 7 |
section\<open>Theory of Events for SET\<close> |
14199 | 8 |
|
33028 | 9 |
theory Event_SET |
10 |
imports Message_SET |
|
11 |
begin |
|
14199 | 12 |
|
63167 | 13 |
text\<open>The Root Certification Authority\<close> |
35101 | 14 |
abbreviation "RCA == CA 0" |
14199 | 15 |
|
16 |
||
63167 | 17 |
text\<open>Message events\<close> |
58310 | 18 |
datatype |
14199 | 19 |
event = Says agent agent msg |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
20 |
| Gets agent msg |
14199 | 21 |
| Notes agent msg |
22 |
||
23 |
||
63167 | 24 |
text\<open>compromised agents: keys known, Notes visible\<close> |
14199 | 25 |
consts bad :: "agent set" |
26 |
||
63167 | 27 |
text\<open>Spy has access to his own key for spoof messages, but RCA is secure\<close> |
14199 | 28 |
specification (bad) |
29 |
Spy_in_bad [iff]: "Spy \<in> bad" |
|
30 |
RCA_not_bad [iff]: "RCA \<notin> bad" |
|
31 |
by (rule exI [of _ "{Spy}"], simp) |
|
32 |
||
33 |
||
63167 | 34 |
subsection\<open>Agents' Knowledge\<close> |
14199 | 35 |
|
36 |
consts (*Initial states of agents -- parameter of the construction*) |
|
67613 | 37 |
initState :: "agent \<Rightarrow> msg set" |
14199 | 38 |
|
39 |
(* Message reception does not extend spy's knowledge because of |
|
40 |
reception invariant enforced by Reception rule in protocol definition*) |
|
67613 | 41 |
primrec knows :: "[agent, event list] \<Rightarrow> msg set" |
39758 | 42 |
where |
43 |
knows_Nil: |
|
44 |
"knows A [] = initState A" |
|
45 |
| knows_Cons: |
|
14199 | 46 |
"knows A (ev # evs) = |
47 |
(if A = Spy then |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
48 |
(case ev of |
67613 | 49 |
Says A' B X \<Rightarrow> insert X (knows Spy evs) |
50 |
| Gets A' X \<Rightarrow> knows Spy evs |
|
51 |
| Notes A' X \<Rightarrow> |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
52 |
if A' \<in> bad then insert X (knows Spy evs) else knows Spy evs) |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
53 |
else |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
54 |
(case ev of |
67613 | 55 |
Says A' B X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
56 |
if A'=A then insert X (knows A evs) else knows A evs |
67613 | 57 |
| Gets A' X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
58 |
if A'=A then insert X (knows A evs) else knows A evs |
67613 | 59 |
| Notes A' X \<Rightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
60 |
if A'=A then insert X (knows A evs) else knows A evs))" |
14199 | 61 |
|
62 |
||
63167 | 63 |
subsection\<open>Used Messages\<close> |
14199 | 64 |
|
67613 | 65 |
primrec used :: "event list \<Rightarrow> msg set" |
39758 | 66 |
where |
14199 | 67 |
(*Set of items that might be visible to somebody: |
39758 | 68 |
complement of the set of fresh items. |
69 |
As above, message reception does extend used items *) |
|
14199 | 70 |
used_Nil: "used [] = (UN B. parts (initState B))" |
39758 | 71 |
| used_Cons: "used (ev # evs) = |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
30549
diff
changeset
|
72 |
(case ev of |
67613 | 73 |
Says A B X \<Rightarrow> parts {X} \<union> (used evs) |
74 |
| Gets A X \<Rightarrow> used evs |
|
75 |
| Notes A X \<Rightarrow> parts {X} \<union> (used evs))" |
|
14199 | 76 |
|
77 |
||
78 |
||
79 |
(* Inserted by default but later removed. This declaration lets the file |
|
80 |
be re-loaded. Addsimps [knows_Cons, used_Nil, *) |
|
81 |
||
82 |
(** Simplifying parts (insert X (knows Spy evs)) |
|
67613 | 83 |
= parts {X} \<union> parts (knows Spy evs) -- since general case loops*) |
14199 | 84 |
|
45605 | 85 |
lemmas parts_insert_knows_A = parts_insert [of _ "knows A evs"] for A evs |
14199 | 86 |
|
87 |
lemma knows_Spy_Says [simp]: |
|
88 |
"knows Spy (Says A B X # evs) = insert X (knows Spy evs)" |
|
89 |
by auto |
|
90 |
||
63167 | 91 |
text\<open>Letting the Spy see "bad" agents' notes avoids redundant case-splits |
69597 | 92 |
on whether \<^term>\<open>A=Spy\<close> and whether \<^term>\<open>A\<in>bad\<close>\<close> |
14199 | 93 |
lemma knows_Spy_Notes [simp]: |
94 |
"knows Spy (Notes A X # evs) = |
|
67613 | 95 |
(if A\<in>bad then insert X (knows Spy evs) else knows Spy evs)" |
14199 | 96 |
apply auto |
97 |
done |
|
98 |
||
99 |
lemma knows_Spy_Gets [simp]: "knows Spy (Gets A X # evs) = knows Spy evs" |
|
100 |
by auto |
|
101 |
||
67613 | 102 |
lemma initState_subset_knows: "initState A \<subseteq> knows A evs" |
14199 | 103 |
apply (induct_tac "evs") |
104 |
apply (auto split: event.split) |
|
105 |
done |
|
106 |
||
107 |
lemma knows_Spy_subset_knows_Spy_Says: |
|
67613 | 108 |
"knows Spy evs \<subseteq> knows Spy (Says A B X # evs)" |
14199 | 109 |
by auto |
110 |
||
111 |
lemma knows_Spy_subset_knows_Spy_Notes: |
|
67613 | 112 |
"knows Spy evs \<subseteq> knows Spy (Notes A X # evs)" |
14199 | 113 |
by auto |
114 |
||
115 |
lemma knows_Spy_subset_knows_Spy_Gets: |
|
67613 | 116 |
"knows Spy evs \<subseteq> knows Spy (Gets A X # evs)" |
14199 | 117 |
by auto |
118 |
||
119 |
(*Spy sees what is sent on the traffic*) |
|
120 |
lemma Says_imp_knows_Spy [rule_format]: |
|
67613 | 121 |
"Says A B X \<in> set evs \<longrightarrow> X \<in> knows Spy evs" |
14199 | 122 |
apply (induct_tac "evs") |
123 |
apply (auto split: event.split) |
|
124 |
done |
|
125 |
||
126 |
(*Use with addSEs to derive contradictions from old Says events containing |
|
127 |
items known to be fresh*) |
|
128 |
lemmas knows_Spy_partsEs = |
|
46471 | 129 |
Says_imp_knows_Spy [THEN parts.Inj, elim_format] |
130 |
parts.Body [elim_format] |
|
14199 | 131 |
|
132 |
||
69597 | 133 |
subsection\<open>The Function \<^term>\<open>used\<close>\<close> |
14199 | 134 |
|
67613 | 135 |
lemma parts_knows_Spy_subset_used: "parts (knows Spy evs) \<subseteq> used evs" |
14199 | 136 |
apply (induct_tac "evs") |
137 |
apply (auto simp add: parts_insert_knows_A split: event.split) |
|
138 |
done |
|
139 |
||
140 |
lemmas usedI = parts_knows_Spy_subset_used [THEN subsetD, intro] |
|
141 |
||
67613 | 142 |
lemma initState_subset_used: "parts (initState B) \<subseteq> used evs" |
14199 | 143 |
apply (induct_tac "evs") |
144 |
apply (auto split: event.split) |
|
145 |
done |
|
146 |
||
147 |
lemmas initState_into_used = initState_subset_used [THEN subsetD] |
|
148 |
||
67613 | 149 |
lemma used_Says [simp]: "used (Says A B X # evs) = parts{X} \<union> used evs" |
14199 | 150 |
by auto |
151 |
||
67613 | 152 |
lemma used_Notes [simp]: "used (Notes A X # evs) = parts{X} \<union> used evs" |
14199 | 153 |
by auto |
154 |
||
155 |
lemma used_Gets [simp]: "used (Gets A X # evs) = used evs" |
|
156 |
by auto |
|
157 |
||
158 |
||
159 |
lemma Notes_imp_parts_subset_used [rule_format]: |
|
67613 | 160 |
"Notes A X \<in> set evs \<longrightarrow> parts {X} \<subseteq> used evs" |
14199 | 161 |
apply (induct_tac "evs") |
55417
01fbfb60c33e
adapted to 'xxx_{case,rec}' renaming, to new theorem names, and to new variable names in theorems
blanchet
parents:
46471
diff
changeset
|
162 |
apply (rename_tac [2] a evs') |
14199 | 163 |
apply (induct_tac [2] "a", auto) |
164 |
done |
|
165 |
||
63167 | 166 |
text\<open>NOTE REMOVAL--laws above are cleaner, as they don't involve "case"\<close> |
14199 | 167 |
declare knows_Cons [simp del] |
168 |
used_Nil [simp del] used_Cons [simp del] |
|
169 |
||
170 |
||
69597 | 171 |
text\<open>For proving theorems of the form \<^term>\<open>X \<notin> analz (knows Spy evs) \<longrightarrow> P\<close> |
14199 | 172 |
New events added by induction to "evs" are discarded. Provided |
173 |
this information isn't needed, the proof will be much shorter, since |
|
69597 | 174 |
it will omit complicated reasoning about \<^term>\<open>analz\<close>.\<close> |
14199 | 175 |
|
176 |
lemmas analz_mono_contra = |
|
177 |
knows_Spy_subset_knows_Spy_Says [THEN analz_mono, THEN contra_subsetD] |
|
178 |
knows_Spy_subset_knows_Spy_Notes [THEN analz_mono, THEN contra_subsetD] |
|
179 |
knows_Spy_subset_knows_Spy_Gets [THEN analz_mono, THEN contra_subsetD] |
|
27225 | 180 |
|
45605 | 181 |
lemmas analz_impI = impI [where P = "Y \<notin> analz (knows Spy evs)"] for Y evs |
27225 | 182 |
|
14199 | 183 |
ML |
63167 | 184 |
\<open> |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
185 |
fun analz_mono_contra_tac ctxt = |
60754 | 186 |
resolve_tac ctxt @{thms analz_impI} THEN' |
59498
50b60f501b05
proper context for resolve_tac, eresolve_tac, dresolve_tac, forward_tac etc.;
wenzelm
parents:
58963
diff
changeset
|
187 |
REPEAT1 o (dresolve_tac ctxt @{thms analz_mono_contra}) |
58963
26bf09b95dda
proper context for assume_tac (atac remains as fall-back without context);
wenzelm
parents:
58889
diff
changeset
|
188 |
THEN' mp_tac ctxt |
63167 | 189 |
\<close> |
14199 | 190 |
|
63167 | 191 |
method_setup analz_mono_contra = \<open> |
192 |
Scan.succeed (fn ctxt => SIMPLE_METHOD (REPEAT_FIRST (analz_mono_contra_tac ctxt)))\<close> |
|
67613 | 193 |
"for proving theorems of the form X \<notin> analz (knows Spy evs) \<longrightarrow> P" |
14199 | 194 |
|
195 |
end |