src/HOL/Library/Quotient.thy
author wenzelm
Sat, 03 Feb 2001 17:43:05 +0100
changeset 11047 10c51288b00d
parent 10681 ec76e17f73c5
child 11099 b301d1f72552
permissions -rw-r--r--
tuned;
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$
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
     3
    Author:     Markus Wenzel, TU Muenchen
10681
wenzelm
parents: 10551
diff changeset
     4
    License:    GPL (GNU GENERAL PUBLIC LICENSE)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     5
*)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     6
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
     7
header {*
10473
4f15b844fea6 separate rules for function/operation definitions;
wenzelm
parents: 10459
diff changeset
     8
  \title{Quotient types}
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
     9
  \author{Markus Wenzel}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    10
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    11
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    12
theory Quotient = Main:
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    13
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    14
text {*
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    15
 We introduce the notion of quotient types over equivalence relations
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    16
 via axiomatic type classes.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    17
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    18
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    19
subsection {* Equivalence relations and quotient types *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    20
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    21
text {*
10390
wenzelm
parents: 10333
diff changeset
    22
 \medskip Type class @{text equiv} models equivalence relations @{text
wenzelm
parents: 10333
diff changeset
    23
 "\<sim> :: 'a => 'a => bool"}.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    24
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    25
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    26
axclass eqv < "term"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    27
consts
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    28
  eqv :: "('a::eqv) => 'a => bool"    (infixl "\<sim>" 50)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    29
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    30
axclass equiv < eqv
10333
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    31
  equiv_refl [intro]: "x \<sim> x"
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    32
  equiv_trans [trans]: "x \<sim> y ==> y \<sim> z ==> x \<sim> z"
f12ff6a4bc7b tuned names;
wenzelm
parents: 10311
diff changeset
    33
  equiv_sym [elim?]: "x \<sim> y ==> y \<sim> x"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    34
10477
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    35
lemma not_equiv_sym [elim?]: "\<not> (x \<sim> y) ==> \<not> (y \<sim> (x::'a::equiv))"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    36
proof -
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    37
  assume "\<not> (x \<sim> y)" thus "\<not> (y \<sim> x)"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    38
    by (rule contrapos_nn) (rule equiv_sym)
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    39
qed
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    40
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    41
lemma not_equiv_trans1 [trans]: "\<not> (x \<sim> y) ==> y \<sim> z ==> \<not> (x \<sim> (z::'a::equiv))"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    42
proof -
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    43
  assume "\<not> (x \<sim> y)" and yz: "y \<sim> z"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    44
  show "\<not> (x \<sim> z)"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    45
  proof
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    46
    assume "x \<sim> z"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    47
    also from yz have "z \<sim> y" ..
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    48
    finally have "x \<sim> y" .
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    49
    thus False by contradiction
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    50
  qed
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    51
qed
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    52
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    53
lemma not_equiv_trans2 [trans]: "x \<sim> y ==> \<not> (y \<sim> z) ==> \<not> (x \<sim> (z::'a::equiv))"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    54
proof -
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    55
  assume "\<not> (y \<sim> z)" hence "\<not> (z \<sim> y)" ..
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    56
  also assume "x \<sim> y" hence "y \<sim> x" ..
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    57
  finally have "\<not> (z \<sim> x)" . thus "(\<not> x \<sim> z)" ..
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    58
qed
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
    59
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    60
text {*
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    61
 \medskip The quotient type @{text "'a quot"} consists of all
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    62
 \emph{equivalence classes} over elements of the base type @{typ 'a}.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    63
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    64
10392
wenzelm
parents: 10390
diff changeset
    65
typedef 'a quot = "{{x. a \<sim> x} | a::'a::eqv. True}"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    66
  by blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    67
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    68
lemma quotI [intro]: "{x. a \<sim> x} \<in> quot"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    69
  by (unfold quot_def) blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    70
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    71
lemma quotE [elim]: "R \<in> quot ==> (!!a. R = {x. a \<sim> x} ==> C) ==> C"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    72
  by (unfold quot_def) blast
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    73
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    74
text {*
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    75
 \medskip Abstracted equivalence classes are the canonical
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    76
 representation of elements of a quotient type.
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    77
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    78
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    79
constdefs
10551
ec9fab41b3a0 renamed "equivalence_class" to "class";
wenzelm
parents: 10505
diff changeset
    80
  class :: "'a::equiv => 'a quot"    ("\<lfloor>_\<rfloor>")
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    81
  "\<lfloor>a\<rfloor> == Abs_quot {x. a \<sim> x}"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    82
10311
wenzelm
parents: 10286
diff changeset
    83
theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>"
10278
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    84
proof (cases A)
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    85
  fix R assume R: "A = Abs_quot R"
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    86
  assume "R \<in> quot" hence "\<exists>a. R = {x. a \<sim> x}" by blast
ea1bf4b6255c improved typedef;
wenzelm
parents: 10250
diff changeset
    87
  with R have "\<exists>a. A = Abs_quot {x. a \<sim> x}" by blast
10551
ec9fab41b3a0 renamed "equivalence_class" to "class";
wenzelm
parents: 10505
diff changeset
    88
  thus ?thesis by (unfold class_def)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    89
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    90
10311
wenzelm
parents: 10286
diff changeset
    91
lemma quot_cases [cases type: quot]: "(!!a. A = \<lfloor>a\<rfloor> ==> C) ==> C"
wenzelm
parents: 10286
diff changeset
    92
  by (insert quot_exhaust) blast
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    93
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    94
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
    95
subsection {* Equality on quotients *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    96
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
    97
text {*
10286
wenzelm
parents: 10285
diff changeset
    98
 Equality of canonical quotient elements coincides with the original
wenzelm
parents: 10285
diff changeset
    99
 relation.
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   100
*}
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   101
10477
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   102
theorem quot_equality: "(\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> b)"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   103
proof
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   104
  assume eq: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   105
  show "a \<sim> b"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   106
  proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   107
    from eq have "{x. a \<sim> x} = {x. b \<sim> x}"
10551
ec9fab41b3a0 renamed "equivalence_class" to "class";
wenzelm
parents: 10505
diff changeset
   108
      by (simp only: class_def Abs_quot_inject quotI)
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   109
    moreover have "a \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   110
    ultimately have "a \<in> {x. b \<sim> x}" by blast
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   111
    hence "b \<sim> a" by blast
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   112
    thus ?thesis ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   113
  qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   114
next
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   115
  assume ab: "a \<sim> b"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   116
  show "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   117
  proof -
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   118
    have "{x. a \<sim> x} = {x. b \<sim> x}"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   119
    proof (rule Collect_cong)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   120
      fix x show "(a \<sim> x) = (b \<sim> x)"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   121
      proof
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   122
        from ab have "b \<sim> a" ..
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   123
        also assume "a \<sim> x"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   124
        finally show "b \<sim> x" .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   125
      next
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   126
        note ab
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   127
        also assume "b \<sim> x"
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   128
        finally show "a \<sim> x" .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   129
      qed
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   130
    qed
10551
ec9fab41b3a0 renamed "equivalence_class" to "class";
wenzelm
parents: 10505
diff changeset
   131
    thus ?thesis by (simp only: class_def)
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   132
  qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   133
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   134
10477
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   135
lemma quot_equalI [intro?]: "a \<sim> b ==> \<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   136
  by (simp only: quot_equality)
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   137
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   138
lemma quot_equalD [dest?]: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor> ==> a \<sim> b"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   139
  by (simp only: quot_equality)
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   140
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   141
lemma quot_not_equalI [intro?]: "\<not> (a \<sim> b) ==> \<lfloor>a\<rfloor> \<noteq> \<lfloor>b\<rfloor>"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   142
  by (simp add: quot_equality)
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   143
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   144
lemma quot_not_equalD [dest?]: "\<lfloor>a\<rfloor> \<noteq> \<lfloor>b\<rfloor> ==> \<not> (a \<sim> b)"
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   145
  by (simp add: quot_equality)
c21bee84cefe added not_equiv_sym, not_equiv_trans1/2;
wenzelm
parents: 10473
diff changeset
   146
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   147
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   148
subsection {* Picking representing elements *}
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   149
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   150
constdefs
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   151
  pick :: "'a::equiv quot => 'a"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   152
  "pick A == SOME a. A = \<lfloor>a\<rfloor>"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   153
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   154
theorem pick_equiv [intro]: "pick \<lfloor>a\<rfloor> \<sim> a"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   155
proof (unfold pick_def)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   156
  show "(SOME x. \<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>) \<sim> a"
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   157
  proof (rule someI2)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   158
    show "\<lfloor>a\<rfloor> = \<lfloor>a\<rfloor>" ..
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   159
    fix x assume "\<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   160
    hence "a \<sim> x" .. thus "x \<sim> a" ..
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   161
  qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   162
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   163
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   164
theorem pick_inverse [intro]: "\<lfloor>pick A\<rfloor> = A"
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   165
proof (cases A)
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   166
  fix a assume a: "A = \<lfloor>a\<rfloor>"
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   167
  hence "pick A \<sim> a" by (simp only: pick_equiv)
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   168
  hence "\<lfloor>pick A\<rfloor> = \<lfloor>a\<rfloor>" ..
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   169
  with a show ?thesis by simp
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   170
qed
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   171
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   172
text {*
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   173
 \medskip The following rules support canonical function definitions
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   174
 on quotient types (with up to two arguments).  Note that the
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   175
 stripped-down version without additional conditions is sufficient
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   176
 most of the time.
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   177
*}
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   178
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   179
theorem quot_cond_function:
10491
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   180
  "(!!X Y. P X Y ==> f X Y == g (pick X) (pick Y)) ==>
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   181
    (!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor>
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   182
      ==> P \<lfloor>x\<rfloor> \<lfloor>y\<rfloor> ==> P \<lfloor>x'\<rfloor> \<lfloor>y'\<rfloor> ==> g x y = g x' y') ==>
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   183
    P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> ==> f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b"
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   184
  (is "PROP ?eq ==> PROP ?cong ==> _ ==> _")
10473
4f15b844fea6 separate rules for function/operation definitions;
wenzelm
parents: 10459
diff changeset
   185
proof -
10491
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   186
  assume cong: "PROP ?cong"
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   187
  assume "PROP ?eq" and "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>"
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   188
  hence "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g (pick \<lfloor>a\<rfloor>) (pick \<lfloor>b\<rfloor>)" by (simp only:)
10505
b89e6cc963e3 unsymbolize;
wenzelm
parents: 10499
diff changeset
   189
  also have "... = g a b"
10491
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   190
  proof (rule cong)
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   191
    show "\<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> = \<lfloor>a\<rfloor>" ..
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   192
    moreover
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   193
    show "\<lfloor>pick \<lfloor>b\<rfloor>\<rfloor> = \<lfloor>b\<rfloor>" ..
10491
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   194
    moreover
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   195
    show "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>" .
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   196
    ultimately show "P \<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> \<lfloor>pick \<lfloor>b\<rfloor>\<rfloor>" by (simp only:)
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   197
  qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   198
  finally show ?thesis .
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   199
qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   200
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   201
theorem quot_function:
10473
4f15b844fea6 separate rules for function/operation definitions;
wenzelm
parents: 10459
diff changeset
   202
  "(!!X Y. f X Y == g (pick X) (pick Y)) ==>
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   203
    (!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor> ==> g x y = g x' y') ==>
10473
4f15b844fea6 separate rules for function/operation definitions;
wenzelm
parents: 10459
diff changeset
   204
    f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b"
4f15b844fea6 separate rules for function/operation definitions;
wenzelm
parents: 10459
diff changeset
   205
proof -
10491
e4a408728012 quot_cond_function: simplified, support conditional definition;
wenzelm
parents: 10483
diff changeset
   206
  case antecedent from this TrueI
10483
eb93ace45a6e removed quot_cond_function1, quot_function1;
wenzelm
parents: 10477
diff changeset
   207
  show ?thesis by (rule quot_cond_function)
10285
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   208
qed
6949e17f314a simplified quotients (only plain total equivs);
wenzelm
parents: 10278
diff changeset
   209
10499
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   210
theorem quot_function':
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   211
  "(!!X Y. f X Y == g (pick X) (pick Y)) ==>
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   212
    (!!x x' y y'. x \<sim> x' ==> y \<sim> y' ==> g x y = g x' y') ==>
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   213
    f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b"
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   214
  by  (rule quot_function) (simp only: quot_equality)+
2f33d0fd242e alternative function definition;
bauerg
parents: 10491
diff changeset
   215
10250
ca93fe25a84b Quotient types;
wenzelm
parents:
diff changeset
   216
end