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