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