author | haftmann |
Tue, 23 Jun 2009 14:24:58 +0200 | |
changeset 31776 | 151c3f5f28f9 |
parent 30738 | 0842e906300c |
permissions | -rw-r--r-- |
10250 | 1 |
(* Title: HOL/Library/Quotient.thy |
10483 | 2 |
Author: Markus Wenzel, TU Muenchen |
10250 | 3 |
*) |
4 |
||
14706 | 5 |
header {* Quotient types *} |
10250 | 6 |
|
15131 | 7 |
theory Quotient |
30738 | 8 |
imports Main |
15131 | 9 |
begin |
10250 | 10 |
|
11 |
text {* |
|
10285 | 12 |
We introduce the notion of quotient types over equivalence relations |
22390 | 13 |
via type classes. |
10250 | 14 |
*} |
15 |
||
10285 | 16 |
subsection {* Equivalence relations and quotient types *} |
10250 | 17 |
|
18 |
text {* |
|
10390 | 19 |
\medskip Type class @{text equiv} models equivalence relations @{text |
20 |
"\<sim> :: 'a => 'a => bool"}. |
|
10250 | 21 |
*} |
22 |
||
29608 | 23 |
class eqv = |
25062 | 24 |
fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sim>" 50) |
10250 | 25 |
|
22390 | 26 |
class equiv = eqv + |
25062 | 27 |
assumes equiv_refl [intro]: "x \<sim> x" |
28 |
assumes equiv_trans [trans]: "x \<sim> y \<Longrightarrow> y \<sim> z \<Longrightarrow> x \<sim> z" |
|
29 |
assumes equiv_sym [sym]: "x \<sim> y \<Longrightarrow> y \<sim> x" |
|
10250 | 30 |
|
12371 | 31 |
lemma equiv_not_sym [sym]: "\<not> (x \<sim> y) ==> \<not> (y \<sim> (x::'a::equiv))" |
10477 | 32 |
proof - |
23373 | 33 |
assume "\<not> (x \<sim> y)" then show "\<not> (y \<sim> x)" |
10477 | 34 |
by (rule contrapos_nn) (rule equiv_sym) |
35 |
qed |
|
36 |
||
37 |
lemma not_equiv_trans1 [trans]: "\<not> (x \<sim> y) ==> y \<sim> z ==> \<not> (x \<sim> (z::'a::equiv))" |
|
38 |
proof - |
|
23373 | 39 |
assume "\<not> (x \<sim> y)" and "y \<sim> z" |
10477 | 40 |
show "\<not> (x \<sim> z)" |
41 |
proof |
|
42 |
assume "x \<sim> z" |
|
23373 | 43 |
also from `y \<sim> z` have "z \<sim> y" .. |
10477 | 44 |
finally have "x \<sim> y" . |
23373 | 45 |
with `\<not> (x \<sim> y)` show False by contradiction |
10477 | 46 |
qed |
47 |
qed |
|
48 |
||
49 |
lemma not_equiv_trans2 [trans]: "x \<sim> y ==> \<not> (y \<sim> z) ==> \<not> (x \<sim> (z::'a::equiv))" |
|
50 |
proof - |
|
23373 | 51 |
assume "\<not> (y \<sim> z)" then have "\<not> (z \<sim> y)" .. |
52 |
also assume "x \<sim> y" then have "y \<sim> x" .. |
|
53 |
finally have "\<not> (z \<sim> x)" . then show "(\<not> x \<sim> z)" .. |
|
10477 | 54 |
qed |
55 |
||
10250 | 56 |
text {* |
10285 | 57 |
\medskip The quotient type @{text "'a quot"} consists of all |
58 |
\emph{equivalence classes} over elements of the base type @{typ 'a}. |
|
10250 | 59 |
*} |
60 |
||
10392 | 61 |
typedef 'a quot = "{{x. a \<sim> x} | a::'a::eqv. True}" |
10250 | 62 |
by blast |
63 |
||
64 |
lemma quotI [intro]: "{x. a \<sim> x} \<in> quot" |
|
18730 | 65 |
unfolding quot_def by blast |
10250 | 66 |
|
67 |
lemma quotE [elim]: "R \<in> quot ==> (!!a. R = {x. a \<sim> x} ==> C) ==> C" |
|
18730 | 68 |
unfolding quot_def by blast |
10250 | 69 |
|
70 |
text {* |
|
71 |
\medskip Abstracted equivalence classes are the canonical |
|
72 |
representation of elements of a quotient type. |
|
73 |
*} |
|
74 |
||
19086 | 75 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19086
diff
changeset
|
76 |
"class" :: "'a::equiv => 'a quot" ("\<lfloor>_\<rfloor>") where |
19086 | 77 |
"\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}" |
10250 | 78 |
|
10311 | 79 |
theorem quot_exhaust: "\<exists>a. A = \<lfloor>a\<rfloor>" |
10278 | 80 |
proof (cases A) |
81 |
fix R assume R: "A = Abs_quot R" |
|
23373 | 82 |
assume "R \<in> quot" then have "\<exists>a. R = {x. a \<sim> x}" by blast |
10278 | 83 |
with R have "\<exists>a. A = Abs_quot {x. a \<sim> x}" by blast |
23373 | 84 |
then show ?thesis unfolding class_def . |
10250 | 85 |
qed |
86 |
||
10311 | 87 |
lemma quot_cases [cases type: quot]: "(!!a. A = \<lfloor>a\<rfloor> ==> C) ==> C" |
18730 | 88 |
using quot_exhaust by blast |
10250 | 89 |
|
90 |
||
10285 | 91 |
subsection {* Equality on quotients *} |
10250 | 92 |
|
93 |
text {* |
|
10286 | 94 |
Equality of canonical quotient elements coincides with the original |
95 |
relation. |
|
10250 | 96 |
*} |
97 |
||
12371 | 98 |
theorem quot_equality [iff?]: "(\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>) = (a \<sim> b)" |
10285 | 99 |
proof |
100 |
assume eq: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>" |
|
101 |
show "a \<sim> b" |
|
102 |
proof - |
|
103 |
from eq have "{x. a \<sim> x} = {x. b \<sim> x}" |
|
10551 | 104 |
by (simp only: class_def Abs_quot_inject quotI) |
10285 | 105 |
moreover have "a \<sim> a" .. |
106 |
ultimately have "a \<in> {x. b \<sim> x}" by blast |
|
23373 | 107 |
then have "b \<sim> a" by blast |
108 |
then show ?thesis .. |
|
10285 | 109 |
qed |
110 |
next |
|
10250 | 111 |
assume ab: "a \<sim> b" |
10285 | 112 |
show "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>" |
113 |
proof - |
|
114 |
have "{x. a \<sim> x} = {x. b \<sim> x}" |
|
115 |
proof (rule Collect_cong) |
|
116 |
fix x show "(a \<sim> x) = (b \<sim> x)" |
|
117 |
proof |
|
118 |
from ab have "b \<sim> a" .. |
|
119 |
also assume "a \<sim> x" |
|
120 |
finally show "b \<sim> x" . |
|
121 |
next |
|
122 |
note ab |
|
123 |
also assume "b \<sim> x" |
|
124 |
finally show "a \<sim> x" . |
|
125 |
qed |
|
10250 | 126 |
qed |
23373 | 127 |
then show ?thesis by (simp only: class_def) |
10250 | 128 |
qed |
129 |
qed |
|
130 |
||
131 |
||
10285 | 132 |
subsection {* Picking representing elements *} |
10250 | 133 |
|
19086 | 134 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
19086
diff
changeset
|
135 |
pick :: "'a::equiv quot => 'a" where |
19086 | 136 |
"pick A = (SOME a. A = \<lfloor>a\<rfloor>)" |
10250 | 137 |
|
10285 | 138 |
theorem pick_equiv [intro]: "pick \<lfloor>a\<rfloor> \<sim> a" |
10250 | 139 |
proof (unfold pick_def) |
140 |
show "(SOME x. \<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>) \<sim> a" |
|
141 |
proof (rule someI2) |
|
142 |
show "\<lfloor>a\<rfloor> = \<lfloor>a\<rfloor>" .. |
|
143 |
fix x assume "\<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>" |
|
23373 | 144 |
then have "a \<sim> x" .. then show "x \<sim> a" .. |
10250 | 145 |
qed |
146 |
qed |
|
147 |
||
10483 | 148 |
theorem pick_inverse [intro]: "\<lfloor>pick A\<rfloor> = A" |
10250 | 149 |
proof (cases A) |
150 |
fix a assume a: "A = \<lfloor>a\<rfloor>" |
|
23373 | 151 |
then have "pick A \<sim> a" by (simp only: pick_equiv) |
152 |
then have "\<lfloor>pick A\<rfloor> = \<lfloor>a\<rfloor>" .. |
|
10250 | 153 |
with a show ?thesis by simp |
154 |
qed |
|
155 |
||
10285 | 156 |
text {* |
157 |
\medskip The following rules support canonical function definitions |
|
10483 | 158 |
on quotient types (with up to two arguments). Note that the |
159 |
stripped-down version without additional conditions is sufficient |
|
160 |
most of the time. |
|
10285 | 161 |
*} |
162 |
||
10483 | 163 |
theorem quot_cond_function: |
18372 | 164 |
assumes eq: "!!X Y. P X Y ==> f X Y == g (pick X) (pick Y)" |
165 |
and cong: "!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor> |
|
166 |
==> P \<lfloor>x\<rfloor> \<lfloor>y\<rfloor> ==> P \<lfloor>x'\<rfloor> \<lfloor>y'\<rfloor> ==> g x y = g x' y'" |
|
167 |
and P: "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>" |
|
168 |
shows "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" |
|
10473
4f15b844fea6
separate rules for function/operation definitions;
wenzelm
parents:
10459
diff
changeset
|
169 |
proof - |
18372 | 170 |
from eq and P have "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g (pick \<lfloor>a\<rfloor>) (pick \<lfloor>b\<rfloor>)" by (simp only:) |
10505 | 171 |
also have "... = g a b" |
10491
e4a408728012
quot_cond_function: simplified, support conditional definition;
wenzelm
parents:
10483
diff
changeset
|
172 |
proof (rule cong) |
10483 | 173 |
show "\<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> = \<lfloor>a\<rfloor>" .. |
174 |
moreover |
|
175 |
show "\<lfloor>pick \<lfloor>b\<rfloor>\<rfloor> = \<lfloor>b\<rfloor>" .. |
|
10491
e4a408728012
quot_cond_function: simplified, support conditional definition;
wenzelm
parents:
10483
diff
changeset
|
176 |
moreover |
23373 | 177 |
show "P \<lfloor>a\<rfloor> \<lfloor>b\<rfloor>" by (rule P) |
10491
e4a408728012
quot_cond_function: simplified, support conditional definition;
wenzelm
parents:
10483
diff
changeset
|
178 |
ultimately show "P \<lfloor>pick \<lfloor>a\<rfloor>\<rfloor> \<lfloor>pick \<lfloor>b\<rfloor>\<rfloor>" by (simp only:) |
10285 | 179 |
qed |
180 |
finally show ?thesis . |
|
181 |
qed |
|
182 |
||
10483 | 183 |
theorem quot_function: |
18372 | 184 |
assumes "!!X Y. f X Y == g (pick X) (pick Y)" |
185 |
and "!!x x' y y'. \<lfloor>x\<rfloor> = \<lfloor>x'\<rfloor> ==> \<lfloor>y\<rfloor> = \<lfloor>y'\<rfloor> ==> g x y = g x' y'" |
|
186 |
shows "f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" |
|
23394 | 187 |
using assms and TrueI |
18372 | 188 |
by (rule quot_cond_function) |
10285 | 189 |
|
10499 | 190 |
theorem quot_function': |
191 |
"(!!X Y. f X Y == g (pick X) (pick Y)) ==> |
|
192 |
(!!x x' y y'. x \<sim> x' ==> y \<sim> y' ==> g x y = g x' y') ==> |
|
193 |
f \<lfloor>a\<rfloor> \<lfloor>b\<rfloor> = g a b" |
|
18372 | 194 |
by (rule quot_function) (simp_all only: quot_equality) |
10499 | 195 |
|
10250 | 196 |
end |