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