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