Quot/Examples/Larry.thy
author Christian Urban <urbanc@in.tum.de>
Thu, 10 Dec 2009 18:28:30 +0100
changeset 700 91b079db7380
child 702 b89b578d15e6
permissions -rw-r--r--
added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
700
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     1
theory Larry
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
imports Main "../QuotMain"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
begin
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
subsection{*Defining the Free Algebra*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     7
datatype
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
  freemsg = NONCE  nat
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     9
        | MPAIR  freemsg freemsg
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    10
        | CRYPT  nat freemsg  
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    11
        | DECRYPT  nat freemsg
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    12
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    13
inductive 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    14
  msgrel::"freemsg \<Rightarrow> freemsg \<Rightarrow> bool" (infixl "\<sim>" 50)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    15
where 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    16
  CD:    "CRYPT K (DECRYPT K X) \<sim> X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    17
| DC:    "DECRYPT K (CRYPT K X) \<sim> X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    18
| NONCE: "NONCE N \<sim> NONCE N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    19
| MPAIR: "\<lbrakk>X \<sim> X'; Y \<sim> Y'\<rbrakk> \<Longrightarrow> MPAIR X Y \<sim> MPAIR X' Y'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    20
| CRYPT: "X \<sim> X' \<Longrightarrow> CRYPT K X \<sim> CRYPT K X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    21
| DECRYPT: "X \<sim> X' \<Longrightarrow> DECRYPT K X \<sim> DECRYPT K X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    22
| SYM:   "X \<sim> Y \<Longrightarrow> Y \<sim> X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    23
| TRANS: "\<lbrakk>X \<sim> Y; Y \<sim> Z\<rbrakk> \<Longrightarrow> X \<sim> Z"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    24
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    25
text{*Proving that it is an equivalence relation*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    26
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    27
lemma msgrel_refl: "X \<sim> X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    28
by (induct X, (blast intro: msgrel.intros)+)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    29
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    30
theorem equiv_msgrel: "equivp msgrel"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    31
proof (rule equivpI)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    32
  show "reflp msgrel" by (simp add: reflp_def msgrel_refl)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    33
  show "symp msgrel" by (simp add: symp_def, blast intro: msgrel.SYM)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
  show "transp msgrel" by (simp add: transp_def, blast intro: msgrel.TRANS)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
qed
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
subsection{*Some Functions on the Free Algebra*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    38
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    39
subsubsection{*The Set of Nonces*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    40
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    41
primrec
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
  freenonces :: "freemsg \<Rightarrow> nat set"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
  "freenonces (NONCE N) = {N}"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
| "freenonces (MPAIR X Y) = freenonces X \<union> freenonces Y"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
| "freenonces (CRYPT K X) = freenonces X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
| "freenonces (DECRYPT K X) = freenonces X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    49
theorem msgrel_imp_eq_freenonces: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    50
  "U \<sim> V \<Longrightarrow> freenonces U = freenonces V"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    51
by (erule msgrel.induct, auto) 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    52
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    53
subsubsection{*The Left Projection*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    54
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    55
text{*A function to return the left part of the top pair in a message.  It will
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    56
be lifted to the initial algrebra, to serve as an example of that process.*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    57
fun
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    58
  freeleft :: "freemsg \<Rightarrow> freemsg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    59
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    60
  "freeleft (NONCE N) = NONCE N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    61
| "freeleft (MPAIR X Y) = X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    62
| "freeleft (CRYPT K X) = freeleft X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    63
| "freeleft (DECRYPT K X) = freeleft X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    64
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    65
text{*This theorem lets us prove that the left function respects the
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    66
equivalence relation.  It also helps us prove that MPair
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    67
  (the abstract constructor) is injective*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
lemma msgrel_imp_eqv_freeleft_aux:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    69
  shows "freeleft U \<sim> freeleft U"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
apply(induct rule: freeleft.induct)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    71
apply(auto intro: msgrel.intros)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    72
done
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    73
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    74
theorem msgrel_imp_eqv_freeleft:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    75
  "U \<sim> V \<Longrightarrow> freeleft U \<sim> freeleft V"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    76
apply(erule msgrel.induct)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    77
apply(auto intro: msgrel.intros msgrel_imp_eqv_freeleft_aux)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    78
done
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    79
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    80
subsubsection{*The Right Projection*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    81
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    82
text{*A function to return the right part of the top pair in a message.*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    83
fun
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    84
  freeright :: "freemsg \<Rightarrow> freemsg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    85
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    86
  "freeright (NONCE N) = NONCE N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    87
| "freeright (MPAIR X Y) = Y"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    88
| "freeright (CRYPT K X) = freeright X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    89
| "freeright (DECRYPT K X) = freeright X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    90
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    91
text{*This theorem lets us prove that the right function respects the
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    92
equivalence relation.  It also helps us prove that MPair
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    93
  (the abstract constructor) is injective*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    94
lemma msgrel_imp_eqv_freeright_aux:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    95
  shows "freeright U \<sim> freeright U"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
apply(induct rule: freeright.induct)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    97
apply(auto intro: msgrel.intros)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
done
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    99
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   100
theorem msgrel_imp_eqv_freeright:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   101
     "U \<sim> V \<Longrightarrow> freeright U \<sim> freeright V"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   102
by (erule msgrel.induct, auto intro: msgrel.intros msgrel_imp_eqv_freeright_aux)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   103
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   104
subsubsection{*The Discriminator for Constructors*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   105
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   106
text{*A function to distinguish nonces, mpairs and encryptions*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   107
fun 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   108
  freediscrim :: "freemsg \<Rightarrow> int"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   109
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   110
   "freediscrim (NONCE N) = 0"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   111
 | "freediscrim (MPAIR X Y) = 1"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   112
 | "freediscrim (CRYPT K X) = freediscrim X + 2"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   113
 | "freediscrim (DECRYPT K X) = freediscrim X - 2"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   114
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   115
text{*This theorem helps us prove @{term "Nonce N \<noteq> MPair X Y"}*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   116
theorem msgrel_imp_eq_freediscrim:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   117
     "U \<sim> V \<Longrightarrow> freediscrim U = freediscrim V"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   118
by (erule msgrel.induct, auto)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   119
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   120
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   121
subsection{*The Initial Algebra: A Quotiented Message Type*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   122
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   123
quotient  msg = freemsg / msgrel
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   124
  by (rule equiv_msgrel)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   125
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   126
text{*The abstract message constructors*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   127
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   128
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   129
  Nonce::"Nonce :: nat \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   130
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   131
  "NONCE"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   132
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   133
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   134
  MPair::"MPair :: msg \<Rightarrow> msg \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   135
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   136
  "MPAIR"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   137
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   138
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   139
  Crypt::"Crypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   140
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   141
  "CRYPT"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   142
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   143
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   144
  Decrypt::"Decrypt :: nat \<Rightarrow> msg \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   145
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   146
  "DECRYPT"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   147
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   148
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   149
  shows "(op = ===> op \<sim> ===> op \<sim>) CRYPT CRYPT"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   150
by (auto intro: CRYPT)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   151
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   152
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   153
  shows "(op = ===> op \<sim> ===> op \<sim>) DECRYPT DECRYPT"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   154
by (auto intro: DECRYPT)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   155
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   156
text{*Establishing these two equations is the point of the whole exercise*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   157
theorem CD_eq [simp]: "Crypt K (Decrypt K X) = X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   158
by (lifting CD)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   159
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   160
theorem DC_eq [simp]: "Decrypt K (Crypt K X) = X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   161
by (lifting DC)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   162
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   163
subsection{*The Abstract Function to Return the Set of Nonces*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   164
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   165
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   166
  nonces :: "nounces:: msg \<Rightarrow> nat set"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   167
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   168
  "freenonces"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   169
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   170
text{*Now prove the four equations for @{term nonces}*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   171
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   172
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   173
  shows "(op \<sim> ===> op =) freenonces freenonces"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   174
by (simp add: msgrel_imp_eq_freenonces)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   175
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   176
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   177
  shows "(op = ===> op \<sim>) NONCE NONCE"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   178
by (simp add: NONCE)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
lemma nonces_Nonce [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
  shows "nonces (Nonce N) = {N}"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   182
by (lifting freenonces.simps(1))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   183
 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   184
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   185
  shows " (op \<sim> ===> op \<sim> ===> op \<sim>) MPAIR MPAIR"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   186
by (simp add: MPAIR)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   187
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   188
lemma nonces_MPair [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   189
  shows "nonces (MPair X Y) = nonces X \<union> nonces Y"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   190
by (lifting freenonces.simps(2))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   191
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   192
lemma nonces_Crypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   193
  shows "nonces (Crypt K X) = nonces X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   194
by (lifting freenonces.simps(3))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   195
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   196
lemma nonces_Decrypt [simp]: "nonces (Decrypt K X) = nonces X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   197
by (lifting freenonces.simps(4))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   198
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   199
subsection{*The Abstract Function to Return the Left Part*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   200
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   201
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   202
  left :: "left:: msg \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   203
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   204
  "freeleft"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   205
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   206
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   207
  shows "(op \<sim> ===> op \<sim>) freeleft freeleft"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   208
by (simp add: msgrel_imp_eqv_freeleft)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   209
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   210
lemma left_Nonce [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   211
  shows "left (Nonce N) = Nonce N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   212
by (lifting freeleft.simps(1))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   213
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   214
lemma left_MPair [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   215
  shows "left (MPair X Y) = X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   216
by (lifting freeleft.simps(2))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   217
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   218
lemma left_Crypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   219
  shows "left (Crypt K X) = left X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   220
by (lifting freeleft.simps(3))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   221
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   222
lemma left_Decrypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   223
  shows "left (Decrypt K X) = left X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   224
by (lifting freeleft.simps(4))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   225
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   226
subsection{*The Abstract Function to Return the Right Part*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   227
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   228
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   229
  right :: "right:: msg \<Rightarrow> msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   230
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   231
  "freeright"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   232
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   233
text{*Now prove the four equations for @{term right}*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   234
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   235
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   236
  shows "(op \<sim> ===> op \<sim>) freeright freeright"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   237
by (simp add: msgrel_imp_eqv_freeright)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   238
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   239
lemma right_Nonce [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   240
  shows "right (Nonce N) = Nonce N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   241
by (lifting freeright.simps(1))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
lemma right_MPair [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   244
  shows "right (MPair X Y) = Y"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   245
by (lifting freeright.simps(2))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   246
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   247
lemma right_Crypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   248
  shows "right (Crypt K X) = right X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   249
by (lifting freeright.simps(3))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   250
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   251
lemma right_Decrypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   252
  shows "right (Decrypt K X) = right X"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
by (lifting freeright.simps(4))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   255
subsection{*Injectivity Properties of Some Constructors*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   256
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   257
lemma NONCE_imp_eq: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   258
  shows "NONCE m \<sim> NONCE n \<Longrightarrow> m = n"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   259
by (drule msgrel_imp_eq_freenonces, simp)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   260
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   261
text{*Can also be proved using the function @{term nonces}*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   262
lemma Nonce_Nonce_eq [iff]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   263
  shows "(Nonce m = Nonce n) = (m = n)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   264
apply(rule iffI)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   265
apply(lifting NONCE_imp_eq)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
apply(simp)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   267
done
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   268
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   269
lemma MPAIR_imp_eqv_left: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   270
  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> X \<sim> X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   271
by (drule msgrel_imp_eqv_freeleft, simp)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   272
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   273
lemma MPair_imp_eq_left: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   274
  assumes eq: "MPair X Y = MPair X' Y'" 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   275
  shows "X = X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   276
using eq by (lifting MPAIR_imp_eqv_left)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   277
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   278
lemma MPAIR_imp_eqv_right: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   279
  shows "MPAIR X Y \<sim> MPAIR X' Y' \<Longrightarrow> Y \<sim> Y'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   280
by (drule msgrel_imp_eqv_freeright, simp)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   281
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   282
lemma MPair_imp_eq_right: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   283
  shows "MPair X Y = MPair X' Y' \<Longrightarrow> Y = Y'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   284
by (lifting  MPAIR_imp_eqv_right)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   285
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   286
theorem MPair_MPair_eq [iff]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   287
  shows "(MPair X Y = MPair X' Y') = (X=X' & Y=Y')" 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   288
by (blast dest: MPair_imp_eq_left MPair_imp_eq_right)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   289
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   290
lemma NONCE_neqv_MPAIR: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   291
  shows "\<not>(NONCE m \<sim> MPAIR X Y)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   292
by (auto dest: msgrel_imp_eq_freediscrim)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   293
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   294
theorem Nonce_neq_MPair [iff]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   295
  shows "Nonce N \<noteq> MPair X Y"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   296
by (lifting NONCE_neqv_MPAIR)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   297
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   298
text{*Example suggested by a referee*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   299
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   300
lemma CRYPT_NONCE_neq_NONCE:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   301
  shows "\<not>(CRYPT K (NONCE M) \<sim> NONCE N)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   302
by (auto dest: msgrel_imp_eq_freediscrim)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   303
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   304
theorem Crypt_Nonce_neq_Nonce: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   305
  shows "Crypt K (Nonce M) \<noteq> Nonce N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   306
by (lifting CRYPT_NONCE_neq_NONCE)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   307
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   308
text{*...and many similar results*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   309
lemma CRYPT2_NONCE_neq_NONCE: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   310
  shows "\<not>(CRYPT K (CRYPT K' (NONCE M)) \<sim> NONCE N)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   311
by (auto dest: msgrel_imp_eq_freediscrim)  
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   312
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   313
theorem Crypt2_Nonce_neq_Nonce: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   314
  shows "Crypt K (Crypt K' (Nonce M)) \<noteq> Nonce N"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   315
by (lifting CRYPT2_NONCE_neq_NONCE) 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   316
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   317
theorem Crypt_Crypt_eq [iff]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   318
  shows "(Crypt K X = Crypt K X') = (X=X')" 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   319
proof
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   320
  assume "Crypt K X = Crypt K X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   321
  hence "Decrypt K (Crypt K X) = Decrypt K (Crypt K X')" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   322
  thus "X = X'" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   323
next
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   324
  assume "X = X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   325
  thus "Crypt K X = Crypt K X'" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   326
qed
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   327
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   328
theorem Decrypt_Decrypt_eq [iff]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   329
  shows "(Decrypt K X = Decrypt K X') = (X=X')" 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   330
proof
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   331
  assume "Decrypt K X = Decrypt K X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   332
  hence "Crypt K (Decrypt K X) = Crypt K (Decrypt K X')" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   333
  thus "X = X'" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   334
next
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   335
  assume "X = X'"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   336
  thus "Decrypt K X = Decrypt K X'" by simp
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   337
qed
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   338
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   339
lemma msg_induct_aux:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   340
  shows "\<lbrakk>\<And>N. P (Nonce N);
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   341
          \<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y);
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   342
          \<And>K X. P X \<Longrightarrow> P (Crypt K X);
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   343
          \<And>K X. P X \<Longrightarrow> P (Decrypt K X)\<rbrakk> \<Longrightarrow> P msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   344
by (lifting freemsg.induct)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   345
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   346
lemma msg_induct [case_names Nonce MPair Crypt Decrypt, cases type: msg]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   347
  assumes N: "\<And>N. P (Nonce N)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   348
      and M: "\<And>X Y. \<lbrakk>P X; P Y\<rbrakk> \<Longrightarrow> P (MPair X Y)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   349
      and C: "\<And>K X. P X \<Longrightarrow> P (Crypt K X)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   350
      and D: "\<And>K X. P X \<Longrightarrow> P (Decrypt K X)"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   351
  shows "P msg"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   352
using N M C D by (blast intro:  msg_induct_aux)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   353
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   354
subsection{*The Abstract Discriminator*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   355
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   356
text{*However, as @{text Crypt_Nonce_neq_Nonce} above illustrates, we don't
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   357
need this function in order to prove discrimination theorems.*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   358
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   359
quotient_def
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   360
  discrim :: "discrim:: msg \<Rightarrow> int"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   361
where
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   362
  "freediscrim"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   363
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   364
text{*Now prove the four equations for @{term discrim}*}
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   365
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   366
lemma [quot_respect]:
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   367
  shows "(op \<sim> ===> op =) freediscrim freediscrim"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   368
by (auto simp add: msgrel_imp_eq_freediscrim)
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   369
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   370
lemma discrim_Nonce [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   371
  shows "discrim (Nonce N) = 0"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   372
by (lifting freediscrim.simps(1))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   373
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   374
lemma discrim_MPair [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   375
  shows "discrim (MPair X Y) = 1"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   376
by (lifting freediscrim.simps(2))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   377
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   378
lemma discrim_Crypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   379
  shows "discrim (Crypt K X) = discrim X + 2"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   380
by (lifting freediscrim.simps(3))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   381
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   382
lemma discrim_Decrypt [simp]: 
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   383
  shows "discrim (Decrypt K X) = discrim X - 2"
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   384
by (lifting freediscrim.simps(4))
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   385
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   386
end
91b079db7380 added Larry's theory; introduced lemma equivpI; added something to the TODO about error messages
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   387