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