| author | paulson | 
| Wed, 19 Mar 2008 18:10:23 +0100 | |
| changeset 26333 | 68e5eee47a45 | 
| parent 23746 | a455e69c31cc | 
| child 30198 | 922f944f03b2 | 
| permissions | -rw-r--r-- | 
| 14527 | 1  | 
(* Title: HOL/Induct/QuoDataType  | 
2  | 
ID: $Id$  | 
|
3  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
|
4  | 
Copyright 2004 University of Cambridge  | 
|
5  | 
||
6  | 
*)  | 
|
7  | 
||
8  | 
header{*Defining an Initial Algebra by Quotienting a Free Algebra*}
 | 
|
9  | 
||
| 16417 | 10  | 
theory QuoDataType imports Main begin  | 
| 14527 | 11  | 
|
12  | 
subsection{*Defining the Free Algebra*}
 | 
|
13  | 
||
14  | 
text{*Messages with encryption and decryption as free constructors.*}
 | 
|
15  | 
datatype  | 
|
16  | 
freemsg = NONCE nat  | 
|
17  | 
| MPAIR freemsg freemsg  | 
|
18  | 
| CRYPT nat freemsg  | 
|
19  | 
| DECRYPT nat freemsg  | 
|
20  | 
||
21  | 
text{*The equivalence relation, which makes encryption and decryption inverses
 | 
|
| 23746 | 22  | 
provided the keys are the same.  | 
| 19736 | 23  | 
|
| 23746 | 24  | 
The first two rules are the desired equations. The next four rules  | 
| 14527 | 25  | 
make the equations applicable to subterms. The last two rules are symmetry  | 
26  | 
and transitivity.*}  | 
|
| 23746 | 27  | 
|
28  | 
inductive_set  | 
|
29  | 
msgrel :: "(freemsg * freemsg) set"  | 
|
30  | 
and msg_rel :: "[freemsg, freemsg] => bool" (infixl "\<sim>" 50)  | 
|
31  | 
where  | 
|
32  | 
"X \<sim> Y == (X,Y) \<in> msgrel"  | 
|
33  | 
| CD: "CRYPT K (DECRYPT K X) \<sim> X"  | 
|
34  | 
| DC: "DECRYPT K (CRYPT K X) \<sim> X"  | 
|
35  | 
| NONCE: "NONCE N \<sim> NONCE N"  | 
|
36  | 
| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"  | 
|
37  | 
| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"  | 
|
38  | 
| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"  | 
|
39  | 
| SYM: "X \<sim> Y \<Longrightarrow> Y \<sim> X"  | 
|
40  | 
| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"  | 
|
| 14527 | 41  | 
|
42  | 
||
43  | 
text{*Proving that it is an equivalence relation*}
 | 
|
44  | 
||
45  | 
lemma msgrel_refl: "X \<sim> X"  | 
|
| 18460 | 46  | 
by (induct X) (blast intro: msgrel.intros)+  | 
| 14527 | 47  | 
|
48  | 
theorem equiv_msgrel: "equiv UNIV msgrel"  | 
|
| 18460 | 49  | 
proof -  | 
50  | 
have "reflexive msgrel" by (simp add: refl_def msgrel_refl)  | 
|
51  | 
moreover have "sym msgrel" by (simp add: sym_def, blast intro: msgrel.SYM)  | 
|
52  | 
moreover have "trans msgrel" by (simp add: trans_def, blast intro: msgrel.TRANS)  | 
|
53  | 
ultimately show ?thesis by (simp add: equiv_def)  | 
|
| 14527 | 54  | 
qed  | 
55  | 
||
56  | 
||
57  | 
subsection{*Some Functions on the Free Algebra*}
 | 
|
58  | 
||
59  | 
subsubsection{*The Set of Nonces*}
 | 
|
60  | 
||
61  | 
text{*A function to return the set of nonces present in a message.  It will
 | 
|
62  | 
be lifted to the initial algrebra, to serve as an example of that process.*}  | 
|
63  | 
consts  | 
|
64  | 
freenonces :: "freemsg \<Rightarrow> nat set"  | 
|
65  | 
||
66  | 
primrec  | 
|
67  | 
   "freenonces (NONCE N) = {N}"
 | 
|
68  | 
"freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"  | 
|
69  | 
"freenonces (CRYPT K X) = freenonces X"  | 
|
70  | 
"freenonces (DECRYPT K X) = freenonces X"  | 
|
71  | 
||
72  | 
text{*This theorem lets us prove that the nonces function respects the
 | 
|
73  | 
equivalence relation. It also helps us prove that Nonce  | 
|
74  | 
(the abstract constructor) is injective*}  | 
|
75  | 
theorem msgrel_imp_eq_freenonces: "U \<sim> V \<Longrightarrow> freenonces U = freenonces V"  | 
|
| 18460 | 76  | 
by (induct set: msgrel) auto  | 
| 14527 | 77  | 
|
78  | 
||
79  | 
subsubsection{*The Left Projection*}
 | 
|
80  | 
||
81  | 
text{*A function to return the left part of the top pair in a message.  It will
 | 
|
82  | 
be lifted to the initial algrebra, to serve as an example of that process.*}  | 
|
| 14658 | 83  | 
consts freeleft :: "freemsg \<Rightarrow> freemsg"  | 
| 14527 | 84  | 
primrec  | 
| 14658 | 85  | 
"freeleft (NONCE N) = NONCE N"  | 
86  | 
"freeleft (MPAIR X Y) = X"  | 
|
87  | 
"freeleft (CRYPT K X) = freeleft X"  | 
|
88  | 
"freeleft (DECRYPT K X) = freeleft X"  | 
|
| 14527 | 89  | 
|
90  | 
text{*This theorem lets us prove that the left function respects the
 | 
|
91  | 
equivalence relation. It also helps us prove that MPair  | 
|
92  | 
(the abstract constructor) is injective*}  | 
|
| 14658 | 93  | 
theorem msgrel_imp_eqv_freeleft:  | 
94  | 
"U \<sim> V \<Longrightarrow> freeleft U \<sim> freeleft V"  | 
|
| 18460 | 95  | 
by (induct set: msgrel) (auto intro: msgrel.intros)  | 
| 14527 | 96  | 
|
97  | 
||
98  | 
subsubsection{*The Right Projection*}
 | 
|
99  | 
||
100  | 
text{*A function to return the right part of the top pair in a message.*}
 | 
|
| 14658 | 101  | 
consts freeright :: "freemsg \<Rightarrow> freemsg"  | 
| 14527 | 102  | 
primrec  | 
| 14658 | 103  | 
"freeright (NONCE N) = NONCE N"  | 
104  | 
"freeright (MPAIR X Y) = Y"  | 
|
105  | 
"freeright (CRYPT K X) = freeright X"  | 
|
106  | 
"freeright (DECRYPT K X) = freeright X"  | 
|
| 14527 | 107  | 
|
108  | 
text{*This theorem lets us prove that the right function respects the
 | 
|
109  | 
equivalence relation. It also helps us prove that MPair  | 
|
110  | 
(the abstract constructor) is injective*}  | 
|
| 14658 | 111  | 
theorem msgrel_imp_eqv_freeright:  | 
112  | 
"U \<sim> V \<Longrightarrow> freeright U \<sim> freeright V"  | 
|
| 18460 | 113  | 
by (induct set: msgrel) (auto intro: msgrel.intros)  | 
| 14527 | 114  | 
|
115  | 
||
| 15152 | 116  | 
subsubsection{*The Discriminator for Constructors*}
 | 
| 14527 | 117  | 
|
| 15152 | 118  | 
text{*A function to distinguish nonces, mpairs and encryptions*}
 | 
119  | 
consts freediscrim :: "freemsg \<Rightarrow> int"  | 
|
| 14527 | 120  | 
primrec  | 
| 15152 | 121  | 
"freediscrim (NONCE N) = 0"  | 
122  | 
"freediscrim (MPAIR X Y) = 1"  | 
|
123  | 
"freediscrim (CRYPT K X) = freediscrim X + 2"  | 
|
124  | 
"freediscrim (DECRYPT K X) = freediscrim X - 2"  | 
|
| 14527 | 125  | 
|
126  | 
text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
 | 
|
| 15152 | 127  | 
theorem msgrel_imp_eq_freediscrim:  | 
128  | 
"U \<sim> V \<Longrightarrow> freediscrim U = freediscrim V"  | 
|
| 18460 | 129  | 
by (induct set: msgrel) auto  | 
| 14527 | 130  | 
|
131  | 
||
132  | 
subsection{*The Initial Algebra: A Quotiented Message Type*}
 | 
|
133  | 
||
134  | 
typedef (Msg) msg = "UNIV//msgrel"  | 
|
| 18460 | 135  | 
by (auto simp add: quotient_def)  | 
| 14527 | 136  | 
|
137  | 
||
138  | 
text{*The abstract message constructors*}
 | 
|
| 19736 | 139  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
140  | 
Nonce :: "nat \<Rightarrow> msg" where  | 
| 19736 | 141  | 
  "Nonce N = Abs_Msg(msgrel``{NONCE N})"
 | 
| 14527 | 142  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
143  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
144  | 
MPair :: "[msg,msg] \<Rightarrow> msg" where  | 
| 19736 | 145  | 
"MPair X Y =  | 
| 15120 | 146  | 
       Abs_Msg (\<Union>U \<in> Rep_Msg X. \<Union>V \<in> Rep_Msg Y. msgrel``{MPAIR U V})"
 | 
| 14527 | 147  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
148  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
149  | 
Crypt :: "[nat,msg] \<Rightarrow> msg" where  | 
| 19736 | 150  | 
"Crypt K X =  | 
| 15120 | 151  | 
       Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{CRYPT K U})"
 | 
| 14527 | 152  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
153  | 
definition  | 
| 
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
154  | 
Decrypt :: "[nat,msg] \<Rightarrow> msg" where  | 
| 19736 | 155  | 
"Decrypt K X =  | 
| 15120 | 156  | 
       Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{DECRYPT K U})"
 | 
| 14527 | 157  | 
|
158  | 
||
159  | 
text{*Reduces equality of equivalence classes to the @{term msgrel} relation:
 | 
|
160  | 
  @{term "(msgrel `` {x} = msgrel `` {y}) = ((x,y) \<in> msgrel)"} *}
 | 
|
161  | 
lemmas equiv_msgrel_iff = eq_equiv_class_iff [OF equiv_msgrel UNIV_I UNIV_I]  | 
|
162  | 
||
163  | 
declare equiv_msgrel_iff [simp]  | 
|
164  | 
||
165  | 
||
166  | 
text{*All equivalence classes belong to set of representatives*}
 | 
|
| 15169 | 167  | 
lemma [simp]: "msgrel``{U} \<in> Msg"
 | 
| 14527 | 168  | 
by (auto simp add: Msg_def quotient_def intro: msgrel_refl)  | 
169  | 
||
170  | 
lemma inj_on_Abs_Msg: "inj_on Abs_Msg Msg"  | 
|
171  | 
apply (rule inj_on_inverseI)  | 
|
172  | 
apply (erule Abs_Msg_inverse)  | 
|
173  | 
done  | 
|
174  | 
||
175  | 
text{*Reduces equality on abstractions to equality on representatives*}
 | 
|
176  | 
declare inj_on_Abs_Msg [THEN inj_on_iff, simp]  | 
|
177  | 
||
178  | 
declare Abs_Msg_inverse [simp]  | 
|
179  | 
||
180  | 
||
181  | 
subsubsection{*Characteristic Equations for the Abstract Constructors*}
 | 
|
182  | 
||
183  | 
lemma MPair: "MPair (Abs_Msg(msgrel``{U})) (Abs_Msg(msgrel``{V})) = 
 | 
|
184  | 
              Abs_Msg (msgrel``{MPAIR U V})"
 | 
|
185  | 
proof -  | 
|
| 15169 | 186  | 
  have "(\<lambda>U V. msgrel `` {MPAIR U V}) respects2 msgrel"
 | 
| 14527 | 187  | 
by (simp add: congruent2_def msgrel.MPAIR)  | 
188  | 
thus ?thesis  | 
|
| 14658 | 189  | 
by (simp add: MPair_def UN_equiv_class2 [OF equiv_msgrel equiv_msgrel])  | 
| 14527 | 190  | 
qed  | 
191  | 
||
192  | 
lemma Crypt: "Crypt K (Abs_Msg(msgrel``{U})) = Abs_Msg (msgrel``{CRYPT K U})"
 | 
|
193  | 
proof -  | 
|
| 15169 | 194  | 
  have "(\<lambda>U. msgrel `` {CRYPT K U}) respects msgrel"
 | 
| 14527 | 195  | 
by (simp add: congruent_def msgrel.CRYPT)  | 
196  | 
thus ?thesis  | 
|
197  | 
by (simp add: Crypt_def UN_equiv_class [OF equiv_msgrel])  | 
|
198  | 
qed  | 
|
199  | 
||
200  | 
lemma Decrypt:  | 
|
201  | 
     "Decrypt K (Abs_Msg(msgrel``{U})) = Abs_Msg (msgrel``{DECRYPT K U})"
 | 
|
202  | 
proof -  | 
|
| 15169 | 203  | 
  have "(\<lambda>U. msgrel `` {DECRYPT K U}) respects msgrel"
 | 
| 14527 | 204  | 
by (simp add: congruent_def msgrel.DECRYPT)  | 
205  | 
thus ?thesis  | 
|
206  | 
by (simp add: Decrypt_def UN_equiv_class [OF equiv_msgrel])  | 
|
207  | 
qed  | 
|
208  | 
||
209  | 
text{*Case analysis on the representation of a msg as an equivalence class.*}
 | 
|
210  | 
lemma eq_Abs_Msg [case_names Abs_Msg, cases type: msg]:  | 
|
211  | 
     "(!!U. z = Abs_Msg(msgrel``{U}) ==> P) ==> P"
 | 
|
212  | 
apply (rule Rep_Msg [of z, unfolded Msg_def, THEN quotientE])  | 
|
213  | 
apply (drule arg_cong [where f=Abs_Msg])  | 
|
214  | 
apply (auto simp add: Rep_Msg_inverse intro: msgrel_refl)  | 
|
215  | 
done  | 
|
216  | 
||
217  | 
text{*Establishing these two equations is the point of the whole exercise*}
 | 
|
| 14533 | 218  | 
theorem CD_eq [simp]: "Crypt K (Decrypt K X) = X"  | 
| 14527 | 219  | 
by (cases X, simp add: Crypt Decrypt CD)  | 
220  | 
||
| 14533 | 221  | 
theorem DC_eq [simp]: "Decrypt K (Crypt K X) = X"  | 
| 14527 | 222  | 
by (cases X, simp add: Crypt Decrypt DC)  | 
223  | 
||
224  | 
||
225  | 
subsection{*The Abstract Function to Return the Set of Nonces*}
 | 
|
226  | 
||
| 19736 | 227  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
228  | 
nonces :: "msg \<Rightarrow> nat set" where  | 
| 19736 | 229  | 
"nonces X = (\<Union>U \<in> Rep_Msg X. freenonces U)"  | 
| 14527 | 230  | 
|
| 15169 | 231  | 
lemma nonces_congruent: "freenonces respects msgrel"  | 
| 14527 | 232  | 
by (simp add: congruent_def msgrel_imp_eq_freenonces)  | 
233  | 
||
234  | 
||
235  | 
text{*Now prove the four equations for @{term nonces}*}
 | 
|
236  | 
||
237  | 
lemma nonces_Nonce [simp]: "nonces (Nonce N) = {N}"
 | 
|
238  | 
by (simp add: nonces_def Nonce_def  | 
|
239  | 
UN_equiv_class [OF equiv_msgrel nonces_congruent])  | 
|
240  | 
||
241  | 
lemma nonces_MPair [simp]: "nonces (MPair X Y) = nonces X \<union> nonces Y"  | 
|
242  | 
apply (cases X, cases Y)  | 
|
243  | 
apply (simp add: nonces_def MPair  | 
|
244  | 
UN_equiv_class [OF equiv_msgrel nonces_congruent])  | 
|
245  | 
done  | 
|
246  | 
||
247  | 
lemma nonces_Crypt [simp]: "nonces (Crypt K X) = nonces X"  | 
|
248  | 
apply (cases X)  | 
|
249  | 
apply (simp add: nonces_def Crypt  | 
|
250  | 
UN_equiv_class [OF equiv_msgrel nonces_congruent])  | 
|
251  | 
done  | 
|
252  | 
||
253  | 
lemma nonces_Decrypt [simp]: "nonces (Decrypt K X) = nonces X"  | 
|
254  | 
apply (cases X)  | 
|
255  | 
apply (simp add: nonces_def Decrypt  | 
|
256  | 
UN_equiv_class [OF equiv_msgrel nonces_congruent])  | 
|
257  | 
done  | 
|
258  | 
||
259  | 
||
260  | 
subsection{*The Abstract Function to Return the Left Part*}
 | 
|
261  | 
||
| 19736 | 262  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
263  | 
left :: "msg \<Rightarrow> msg" where  | 
| 19736 | 264  | 
   "left X = Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{freeleft U})"
 | 
| 14527 | 265  | 
|
| 15169 | 266  | 
lemma left_congruent: "(\<lambda>U. msgrel `` {freeleft U}) respects msgrel"
 | 
| 14658 | 267  | 
by (simp add: congruent_def msgrel_imp_eqv_freeleft)  | 
| 14527 | 268  | 
|
269  | 
text{*Now prove the four equations for @{term left}*}
 | 
|
270  | 
||
271  | 
lemma left_Nonce [simp]: "left (Nonce N) = Nonce N"  | 
|
272  | 
by (simp add: left_def Nonce_def  | 
|
273  | 
UN_equiv_class [OF equiv_msgrel left_congruent])  | 
|
274  | 
||
275  | 
lemma left_MPair [simp]: "left (MPair X Y) = X"  | 
|
276  | 
apply (cases X, cases Y)  | 
|
277  | 
apply (simp add: left_def MPair  | 
|
278  | 
UN_equiv_class [OF equiv_msgrel left_congruent])  | 
|
279  | 
done  | 
|
280  | 
||
281  | 
lemma left_Crypt [simp]: "left (Crypt K X) = left X"  | 
|
282  | 
apply (cases X)  | 
|
283  | 
apply (simp add: left_def Crypt  | 
|
284  | 
UN_equiv_class [OF equiv_msgrel left_congruent])  | 
|
285  | 
done  | 
|
286  | 
||
287  | 
lemma left_Decrypt [simp]: "left (Decrypt K X) = left X"  | 
|
288  | 
apply (cases X)  | 
|
289  | 
apply (simp add: left_def Decrypt  | 
|
290  | 
UN_equiv_class [OF equiv_msgrel left_congruent])  | 
|
291  | 
done  | 
|
292  | 
||
293  | 
||
294  | 
subsection{*The Abstract Function to Return the Right Part*}
 | 
|
295  | 
||
| 19736 | 296  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
297  | 
right :: "msg \<Rightarrow> msg" where  | 
| 19736 | 298  | 
   "right X = Abs_Msg (\<Union>U \<in> Rep_Msg X. msgrel``{freeright U})"
 | 
| 14527 | 299  | 
|
| 15169 | 300  | 
lemma right_congruent: "(\<lambda>U. msgrel `` {freeright U}) respects msgrel"
 | 
| 14658 | 301  | 
by (simp add: congruent_def msgrel_imp_eqv_freeright)  | 
| 14527 | 302  | 
|
303  | 
text{*Now prove the four equations for @{term right}*}
 | 
|
304  | 
||
305  | 
lemma right_Nonce [simp]: "right (Nonce N) = Nonce N"  | 
|
306  | 
by (simp add: right_def Nonce_def  | 
|
307  | 
UN_equiv_class [OF equiv_msgrel right_congruent])  | 
|
308  | 
||
309  | 
lemma right_MPair [simp]: "right (MPair X Y) = Y"  | 
|
310  | 
apply (cases X, cases Y)  | 
|
311  | 
apply (simp add: right_def MPair  | 
|
312  | 
UN_equiv_class [OF equiv_msgrel right_congruent])  | 
|
313  | 
done  | 
|
314  | 
||
315  | 
lemma right_Crypt [simp]: "right (Crypt K X) = right X"  | 
|
316  | 
apply (cases X)  | 
|
317  | 
apply (simp add: right_def Crypt  | 
|
318  | 
UN_equiv_class [OF equiv_msgrel right_congruent])  | 
|
319  | 
done  | 
|
320  | 
||
321  | 
lemma right_Decrypt [simp]: "right (Decrypt K X) = right X"  | 
|
322  | 
apply (cases X)  | 
|
323  | 
apply (simp add: right_def Decrypt  | 
|
324  | 
UN_equiv_class [OF equiv_msgrel right_congruent])  | 
|
325  | 
done  | 
|
326  | 
||
327  | 
||
328  | 
subsection{*Injectivity Properties of Some Constructors*}
 | 
|
329  | 
||
330  | 
lemma NONCE_imp_eq: "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"  | 
|
331  | 
by (drule msgrel_imp_eq_freenonces, simp)  | 
|
332  | 
||
333  | 
text{*Can also be proved using the function @{term nonces}*}
 | 
|
334  | 
lemma Nonce_Nonce_eq [iff]: "(Nonce m = Nonce n) = (m = n)"  | 
|
335  | 
by (auto simp add: Nonce_def msgrel_refl dest: NONCE_imp_eq)  | 
|
336  | 
||
337  | 
lemma MPAIR_imp_eqv_left: "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"  | 
|
| 14658 | 338  | 
by (drule msgrel_imp_eqv_freeleft, simp)  | 
| 14527 | 339  | 
|
340  | 
lemma MPair_imp_eq_left:  | 
|
341  | 
assumes eq: "MPair X Y = MPair X' Y'" shows "X = X'"  | 
|
342  | 
proof -  | 
|
343  | 
from eq  | 
|
344  | 
have "left (MPair X Y) = left (MPair X' Y')" by simp  | 
|
345  | 
thus ?thesis by simp  | 
|
346  | 
qed  | 
|
347  | 
||
348  | 
lemma MPAIR_imp_eqv_right: "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"  | 
|
| 14658 | 349  | 
by (drule msgrel_imp_eqv_freeright, simp)  | 
| 14527 | 350  | 
|
351  | 
lemma MPair_imp_eq_right: "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"  | 
|
352  | 
apply (cases X, cases X', cases Y, cases Y')  | 
|
353  | 
apply (simp add: MPair)  | 
|
354  | 
apply (erule MPAIR_imp_eqv_right)  | 
|
355  | 
done  | 
|
356  | 
||
357  | 
theorem MPair_MPair_eq [iff]: "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')"  | 
|
| 14533 | 358  | 
by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)  | 
| 14527 | 359  | 
|
360  | 
lemma NONCE_neqv_MPAIR: "NONCE m \<sim> MPAIR X Y \<Longrightarrow> False"  | 
|
| 15152 | 361  | 
by (drule msgrel_imp_eq_freediscrim, simp)  | 
| 14527 | 362  | 
|
363  | 
theorem Nonce_neq_MPair [iff]: "Nonce N \<noteq> MPair X Y"  | 
|
364  | 
apply (cases X, cases Y)  | 
|
365  | 
apply (simp add: Nonce_def MPair)  | 
|
366  | 
apply (blast dest: NONCE_neqv_MPAIR)  | 
|
367  | 
done  | 
|
368  | 
||
| 15152 | 369  | 
text{*Example suggested by a referee*}
 | 
370  | 
theorem Crypt_Nonce_neq_Nonce: "Crypt K (Nonce M) \<noteq> Nonce N"  | 
|
371  | 
by (auto simp add: Nonce_def Crypt dest: msgrel_imp_eq_freediscrim)  | 
|
372  | 
||
373  | 
text{*...and many similar results*}
 | 
|
| 15172 | 374  | 
theorem Crypt2_Nonce_neq_Nonce: "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"  | 
| 15152 | 375  | 
by (auto simp add: Nonce_def Crypt dest: msgrel_imp_eq_freediscrim)  | 
376  | 
||
| 14533 | 377  | 
theorem Crypt_Crypt_eq [iff]: "(Crypt K X = Crypt K X') = (X=X')"  | 
378  | 
proof  | 
|
379  | 
assume "Crypt K X = Crypt K X'"  | 
|
380  | 
hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp  | 
|
381  | 
thus "X = X'" by simp  | 
|
382  | 
next  | 
|
383  | 
assume "X = X'"  | 
|
384  | 
thus "Crypt K X = Crypt K X'" by simp  | 
|
385  | 
qed  | 
|
386  | 
||
387  | 
theorem Decrypt_Decrypt_eq [iff]: "(Decrypt K X = Decrypt K X') = (X=X')"  | 
|
388  | 
proof  | 
|
389  | 
assume "Decrypt K X = Decrypt K X'"  | 
|
390  | 
hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp  | 
|
391  | 
thus "X = X'" by simp  | 
|
392  | 
next  | 
|
393  | 
assume "X = X'"  | 
|
394  | 
thus "Decrypt K X = Decrypt K X'" by simp  | 
|
395  | 
qed  | 
|
396  | 
||
397  | 
lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:  | 
|
398  | 
assumes N: "\<And>N. P (Nonce N)"  | 
|
399  | 
and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"  | 
|
400  | 
and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"  | 
|
401  | 
and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"  | 
|
402  | 
shows "P msg"  | 
|
| 18460 | 403  | 
proof (cases msg)  | 
404  | 
case (Abs_Msg U)  | 
|
405  | 
  have "P (Abs_Msg (msgrel `` {U}))"
 | 
|
| 14533 | 406  | 
proof (induct U)  | 
407  | 
case (NONCE N)  | 
|
408  | 
with N show ?case by (simp add: Nonce_def)  | 
|
409  | 
next  | 
|
410  | 
case (MPAIR X Y)  | 
|
411  | 
    with M [of "Abs_Msg (msgrel `` {X})" "Abs_Msg (msgrel `` {Y})"]
 | 
|
412  | 
show ?case by (simp add: MPair)  | 
|
413  | 
next  | 
|
414  | 
case (CRYPT K X)  | 
|
415  | 
    with C [of "Abs_Msg (msgrel `` {X})"]
 | 
|
416  | 
show ?case by (simp add: Crypt)  | 
|
417  | 
next  | 
|
418  | 
case (DECRYPT K X)  | 
|
419  | 
    with D [of "Abs_Msg (msgrel `` {X})"]
 | 
|
420  | 
show ?case by (simp add: Decrypt)  | 
|
421  | 
qed  | 
|
| 18460 | 422  | 
with Abs_Msg show ?thesis by (simp only:)  | 
| 14533 | 423  | 
qed  | 
424  | 
||
| 15152 | 425  | 
|
426  | 
subsection{*The Abstract Discriminator*}
 | 
|
427  | 
||
428  | 
text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
 | 
|
429  | 
need this function in order to prove discrimination theorems.*}  | 
|
430  | 
||
| 19736 | 431  | 
definition  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21210 
diff
changeset
 | 
432  | 
discrim :: "msg \<Rightarrow> int" where  | 
| 19736 | 433  | 
   "discrim X = contents (\<Union>U \<in> Rep_Msg X. {freediscrim U})"
 | 
| 15152 | 434  | 
|
| 15169 | 435  | 
lemma discrim_congruent: "(\<lambda>U. {freediscrim U}) respects msgrel"
 | 
| 15152 | 436  | 
by (simp add: congruent_def msgrel_imp_eq_freediscrim)  | 
437  | 
||
438  | 
text{*Now prove the four equations for @{term discrim}*}
 | 
|
439  | 
||
440  | 
lemma discrim_Nonce [simp]: "discrim (Nonce N) = 0"  | 
|
441  | 
by (simp add: discrim_def Nonce_def  | 
|
442  | 
UN_equiv_class [OF equiv_msgrel discrim_congruent])  | 
|
443  | 
||
444  | 
lemma discrim_MPair [simp]: "discrim (MPair X Y) = 1"  | 
|
445  | 
apply (cases X, cases Y)  | 
|
446  | 
apply (simp add: discrim_def MPair  | 
|
447  | 
UN_equiv_class [OF equiv_msgrel discrim_congruent])  | 
|
448  | 
done  | 
|
449  | 
||
450  | 
lemma discrim_Crypt [simp]: "discrim (Crypt K X) = discrim X + 2"  | 
|
451  | 
apply (cases X)  | 
|
452  | 
apply (simp add: discrim_def Crypt  | 
|
453  | 
UN_equiv_class [OF equiv_msgrel discrim_congruent])  | 
|
454  | 
done  | 
|
455  | 
||
456  | 
lemma discrim_Decrypt [simp]: "discrim (Decrypt K X) = discrim X - 2"  | 
|
457  | 
apply (cases X)  | 
|
458  | 
apply (simp add: discrim_def Decrypt  | 
|
459  | 
UN_equiv_class [OF equiv_msgrel discrim_congruent])  | 
|
460  | 
done  | 
|
461  | 
||
462  | 
||
| 14527 | 463  | 
end  | 
464  |