src/HOL/Library/Quotient.thy
author wenzelm
Fri, 03 Nov 2000 21:35:36 +0100
changeset 10389 c7d8901ab269
parent 10333 f12ff6a4bc7b
child 10390 1d54567bed24
permissions -rw-r--r--
proper setup of "parallel"; removed unused rules;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     1
(*  Title:      HOL/Library/Quotient.thy
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     2
    ID:         $Id$
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     3
    Author:     Gertrud Bauer and Markus Wenzel, TU Muenchen
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     4
*)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     5
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     6
header {*
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     7
  \title{Quotients}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     8
  \author{Gertrud Bauer and Markus Wenzel}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     9
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    10
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    11
theory Quotient = Main:
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    12
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    13
text {*
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    14
 We introduce the notion of quotient types over equivalence relations
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    15
 via axiomatic type classes.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    16
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    17
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    18
subsection {* Equivalence relations and quotient types *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    19
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    20
text {*
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    21
 \medskip Type class @{text equiv} models equivalence relations using
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    22
 the polymorphic @{text "\<sim> :: 'a => 'a => bool"} relation.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    23
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    24
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    25
axclass eqv < "term"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    26
consts
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    27
  eqv :: "('a::eqv) => 'a => bool"    (infixl "\<sim>" 50)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    28
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    29
axclass equiv < eqv
10333
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    30
  equiv_refl [intro]: "x \<sim> x"
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    31
  equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    32
  equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    33
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    34
text {*
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    35
 \medskip The quotient type @{text "'a quot"} consists of all
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    36
 \emph{equivalence classes} over elements of the base type @{typ 'a}.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    37
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    38
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    39
typedef 'a quot = "{{x. a \<sim> x}| a::'a::eqv. True}"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    40
  by blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    41
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    42
lemma quotI [intro]: "{x. a \<sim> x} \<in> quot"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    43
  by (unfold quot_def) blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    44
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    45
lemma quotE [elim]: "R \<in> quot ==> (!!a. R = {x. a \<sim> x} ==> C) ==> C"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    46
  by (unfold quot_def) blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    47
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    48
text {*
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    49
 \medskip Abstracted equivalence classes are the canonical
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    50
 representation of elements of a quotient type.
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    51
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    52
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    53
constdefs
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    54
  equivalence_class :: "'a::equiv => 'a quot"    ("\<lfloor>_\<rfloor>")
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    55
  "\<lfloor>a\<rfloor> == Abs_quot {x. a \<sim> x}"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    56
10311
wenzelm
parents: 10286
diff changeset
    57
theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>"
10278
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    58
proof (cases A)
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    59
  fix R assume R: "A = Abs_quot R"
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    60
  assume "R \<in> quot" hence "\<exists>a. R = {x. a \<sim> x}" by blast
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    61
  with R have "\<exists>a. A = Abs_quot {x. a \<sim> x}" by blast
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    62
  thus ?thesis by (unfold equivalence_class_def)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    63
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    64
10311
wenzelm
parents: 10286
diff changeset
    65
lemma quot_cases [cases type: quot]: "(!!a. A = \<lfloor>a\<rfloor> ==> C) ==> C"
wenzelm
parents: 10286
diff changeset
    66
  by (insert quot_exhaust) blast
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    67
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    68
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    69
subsection {* Equality on quotients *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    70
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    71
text {*
10286
wenzelm
parents: 10285
diff changeset
    72
 Equality of canonical quotient elements coincides with the original
wenzelm
parents: 10285
diff changeset
    73
 relation.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    74
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    75
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    76
theorem equivalence_class_eq [iff?]: "(\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> b)"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    77
proof
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    78
  assume eq: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    79
  show "a \<sim> b"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    80
  proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    81
    from eq have "{x. a \<sim> x} = {x. b \<sim> x}"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    82
      by (simp only: equivalence_class_def Abs_quot_inject quotI)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    83
    moreover have "a \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    84
    ultimately have "a \<in> {x. b \<sim> x}" by blast
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    85
    hence "b \<sim> a" by blast
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    86
    thus ?thesis ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    87
  qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    88
next
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    89
  assume ab: "a \<sim> b"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    90
  show "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    91
  proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    92
    have "{x. a \<sim> x} = {x. b \<sim> x}"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    93
    proof (rule Collect_cong)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    94
      fix x show "(a \<sim> x) = (b \<sim> x)"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    95
      proof
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    96
        from ab have "b \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    97
        also assume "a \<sim> x"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    98
        finally show "b \<sim> x" .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    99
      next
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   100
        note ab
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   101
        also assume "b \<sim> x"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   102
        finally show "a \<sim> x" .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   103
      qed
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   104
    qed
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   105
    thus ?thesis by (simp only: equivalence_class_def)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   106
  qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   107
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   108
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   109
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   110
subsection {* Picking representing elements *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   111
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   112
constdefs
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   113
  pick :: "'a::equiv quot => 'a"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   114
  "pick A == SOME a. A = \<lfloor>a\<rfloor>"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   115
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   116
theorem pick_equiv [intro]: "pick \<lfloor>a\<rfloor> \<sim> a"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   117
proof (unfold pick_def)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   118
  show "(SOME x. \<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>) \<sim> a"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   119
  proof (rule someI2)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   120
    show "\<lfloor>a\<rfloor> = \<lfloor>a\<rfloor>" ..
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   121
    fix x assume "\<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   122
    hence "a \<sim> x" .. thus "x \<sim> a" ..
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   123
  qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   124
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   125
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   126
theorem pick_inverse: "\<lfloor>pick A\<rfloor> = A"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   127
proof (cases A)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   128
  fix a assume a: "A = \<lfloor>a\<rfloor>"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   129
  hence "pick A \<sim> a" by (simp only: pick_equiv)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   130
  hence "\<lfloor>pick A\<rfloor> = \<lfloor>a\<rfloor>" ..
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   131
  with a show ?thesis by simp
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   132
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   133
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   134
text {*
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   135
 \medskip The following rules support canonical function definitions
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   136
 on quotient types.
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   137
*}
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   138
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   139
theorem cong_definition1:
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   140
  "(!!X. f X == g (pick X)) ==>
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   141
    (!!x x'. x \<sim> x' ==> g x = g x') ==>
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   142
    f \<lfloor>a\<rfloor> = g a"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   143
proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   144
  assume cong: "!!x x'. x \<sim> x' ==> g x = g x'"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   145
  assume "!!X. f X == g (pick X)"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   146
  hence "f \<lfloor>a\<rfloor> = g (pick \<lfloor>a\<rfloor>)" by (simp only:)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   147
  also have "\<dots> = g a"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   148
  proof (rule cong)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   149
    show "pick \<lfloor>a\<rfloor> \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   150
  qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   151
  finally show ?thesis .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   152
qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   153
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   154
theorem cong_definition2:
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   155
  "(!!X Y. f X Y == g (pick X) (pick Y)) ==>
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   156
    (!!x x' y y'. x \<sim> x' ==> y \<sim> y' ==> g x y = g x' y') ==>
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   157
    f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   158
proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   159
  assume cong: "!!x x' y y'. x \<sim> x' ==> y \<sim> y' ==> g x y = g x' y'"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   160
  assume "!!X Y. f X Y == g (pick X) (pick Y)"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   161
  hence "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g (pick \<lfloor>a\<rfloor>) (pick \<lfloor>b\<rfloor>)" by (simp only:)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   162
  also have "\<dots> = g a b"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   163
  proof (rule cong)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   164
    show "pick \<lfloor>a\<rfloor> \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   165
    show "pick \<lfloor>b\<rfloor> \<sim> b" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   166
  qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   167
  finally show ?thesis .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   168
qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   169
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   170
end