author | wenzelm |
Tue, 27 Jul 2021 15:31:54 +0200 | |
changeset 74072 | dc98bb7a439b |
parent 69597 | ff784d5a5bfb |
child 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
1 |
(* Title: HOL/ex/PER.thy |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
2 |
Author: Oscar Slotosch and Markus Wenzel, TU Muenchen |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
3 |
*) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
4 |
|
59031 | 5 |
section \<open>Partial equivalence relations\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
6 |
|
59031 | 7 |
theory PER |
8 |
imports Main |
|
9 |
begin |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
10 |
|
59031 | 11 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
12 |
Higher-order quotients are defined over partial equivalence |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
13 |
relations (PERs) instead of total ones. We provide axiomatic type |
61933 | 14 |
classes \<open>equiv < partial_equiv\<close> and a type constructor |
15 |
\<open>'a quot\<close> with basic operations. This development is based |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
16 |
on: |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
17 |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
18 |
Oscar Slotosch: \emph{Higher Order Quotients and their |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
19 |
Implementation in Isabelle HOL.} Elsa L. Gunter and Amy Felty, |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
20 |
editors, Theorem Proving in Higher Order Logics: TPHOLs '97, |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
21 |
Springer LNCS 1275, 1997. |
59031 | 22 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
23 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
24 |
|
59031 | 25 |
subsection \<open>Partial equivalence\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
26 |
|
59031 | 27 |
text \<open> |
61933 | 28 |
Type class \<open>partial_equiv\<close> models partial equivalence |
29 |
relations (PERs) using the polymorphic \<open>\<sim> :: 'a \<Rightarrow> 'a \<Rightarrow> |
|
30 |
bool\<close> relation, which is required to be symmetric and transitive, |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
31 |
but not necessarily reflexive. |
59031 | 32 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
33 |
|
35315 | 34 |
class partial_equiv = |
59031 | 35 |
fixes eqv :: "'a \<Rightarrow> 'a \<Rightarrow> bool" (infixl "\<sim>" 50) |
36 |
assumes partial_equiv_sym [elim?]: "x \<sim> y \<Longrightarrow> y \<sim> x" |
|
37 |
assumes partial_equiv_trans [trans]: "x \<sim> y \<Longrightarrow> y \<sim> z \<Longrightarrow> x \<sim> z" |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
38 |
|
59031 | 39 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
40 |
\medskip The domain of a partial equivalence relation is the set of |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
41 |
reflexive elements. Due to symmetry and transitivity this |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
42 |
characterizes exactly those elements that are connected with |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
43 |
\emph{any} other one. |
59031 | 44 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
45 |
|
19736 | 46 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
20811
diff
changeset
|
47 |
"domain" :: "'a::partial_equiv set" where |
19736 | 48 |
"domain = {x. x \<sim> x}" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
49 |
|
59031 | 50 |
lemma domainI [intro]: "x \<sim> x \<Longrightarrow> x \<in> domain" |
20811 | 51 |
unfolding domain_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
52 |
|
59031 | 53 |
lemma domainD [dest]: "x \<in> domain \<Longrightarrow> x \<sim> x" |
20811 | 54 |
unfolding domain_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
55 |
|
59031 | 56 |
theorem domainI' [elim?]: "x \<sim> y \<Longrightarrow> x \<in> domain" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
57 |
proof |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
58 |
assume xy: "x \<sim> y" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
59 |
also from xy have "y \<sim> x" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
60 |
finally show "x \<sim> x" . |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
61 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
62 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
63 |
|
59031 | 64 |
subsection \<open>Equivalence on function spaces\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
65 |
|
59031 | 66 |
text \<open> |
61933 | 67 |
The \<open>\<sim>\<close> relation is lifted to function spaces. It is |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
68 |
important to note that this is \emph{not} the direct product, but a |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
69 |
structural one corresponding to the congruence property. |
59031 | 70 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
71 |
|
35315 | 72 |
instantiation "fun" :: (partial_equiv, partial_equiv) partial_equiv |
73 |
begin |
|
74 |
||
59031 | 75 |
definition "f \<sim> g \<longleftrightarrow> (\<forall>x \<in> domain. \<forall>y \<in> domain. x \<sim> y \<longrightarrow> f x \<sim> g y)" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
76 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
77 |
lemma partial_equiv_funI [intro?]: |
59031 | 78 |
"(\<And>x y. x \<in> domain \<Longrightarrow> y \<in> domain \<Longrightarrow> x \<sim> y \<Longrightarrow> f x \<sim> g y) \<Longrightarrow> f \<sim> g" |
20811 | 79 |
unfolding eqv_fun_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
80 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
81 |
lemma partial_equiv_funD [dest?]: |
59031 | 82 |
"f \<sim> g \<Longrightarrow> x \<in> domain \<Longrightarrow> y \<in> domain \<Longrightarrow> x \<sim> y \<Longrightarrow> f x \<sim> g y" |
20811 | 83 |
unfolding eqv_fun_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
84 |
|
59031 | 85 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
86 |
The class of partial equivalence relations is closed under function |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
87 |
spaces (in \emph{both} argument positions). |
59031 | 88 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
89 |
|
35315 | 90 |
instance proof |
59031 | 91 |
fix f g h :: "'a::partial_equiv \<Rightarrow> 'b::partial_equiv" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
92 |
assume fg: "f \<sim> g" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
93 |
show "g \<sim> f" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
94 |
proof |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
95 |
fix x y :: 'a |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
96 |
assume x: "x \<in> domain" and y: "y \<in> domain" |
20811 | 97 |
assume "x \<sim> y" then have "y \<sim> x" .. |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
98 |
with fg y x have "f y \<sim> g x" .. |
20811 | 99 |
then show "g x \<sim> f y" .. |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
100 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
101 |
assume gh: "g \<sim> h" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
102 |
show "f \<sim> h" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
103 |
proof |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
104 |
fix x y :: 'a |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
105 |
assume x: "x \<in> domain" and y: "y \<in> domain" and "x \<sim> y" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
106 |
with fg have "f x \<sim> g y" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
107 |
also from y have "y \<sim> y" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
108 |
with gh y y have "g y \<sim> h y" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
109 |
finally show "f x \<sim> h y" . |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
110 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
111 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
112 |
|
35315 | 113 |
end |
114 |
||
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
115 |
|
59031 | 116 |
subsection \<open>Total equivalence\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
117 |
|
59031 | 118 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
119 |
The class of total equivalence relations on top of PERs. It |
61933 | 120 |
coincides with the standard notion of equivalence, i.e.\ \<open>\<sim> |
121 |
:: 'a \<Rightarrow> 'a \<Rightarrow> bool\<close> is required to be reflexive, transitive and |
|
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
122 |
symmetric. |
59031 | 123 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
124 |
|
35315 | 125 |
class equiv = |
126 |
assumes eqv_refl [intro]: "x \<sim> x" |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
127 |
|
59031 | 128 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
129 |
On total equivalences all elements are reflexive, and congruence |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
130 |
holds unconditionally. |
59031 | 131 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
132 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
133 |
theorem equiv_domain [intro]: "(x::'a::equiv) \<in> domain" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
134 |
proof |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
135 |
show "x \<sim> x" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
136 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
137 |
|
59031 | 138 |
theorem equiv_cong [dest?]: "f \<sim> g \<Longrightarrow> x \<sim> y \<Longrightarrow> f x \<sim> g (y::'a::equiv)" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
139 |
proof - |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
140 |
assume "f \<sim> g" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
141 |
moreover have "x \<in> domain" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
142 |
moreover have "y \<in> domain" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
143 |
moreover assume "x \<sim> y" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
144 |
ultimately show ?thesis .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
145 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
146 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
147 |
|
59031 | 148 |
subsection \<open>Quotient types\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
149 |
|
59031 | 150 |
text \<open> |
61933 | 151 |
The quotient type \<open>'a quot\<close> consists of all |
69597 | 152 |
\emph{equivalence classes} over elements of the base type \<^typ>\<open>'a\<close>. |
59031 | 153 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
154 |
|
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
35315
diff
changeset
|
155 |
definition "quot = {{x. a \<sim> x}| a::'a::partial_equiv. True}" |
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
35315
diff
changeset
|
156 |
|
61260 | 157 |
typedef (overloaded) 'a quot = "quot :: 'a::partial_equiv set set" |
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
35315
diff
changeset
|
158 |
unfolding quot_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
159 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
160 |
lemma quotI [intro]: "{x. a \<sim> x} \<in> quot" |
20811 | 161 |
unfolding quot_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
162 |
|
59031 | 163 |
lemma quotE [elim]: "R \<in> quot \<Longrightarrow> (\<And>a. R = {x. a \<sim> x} \<Longrightarrow> C) \<Longrightarrow> C" |
20811 | 164 |
unfolding quot_def by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
165 |
|
59031 | 166 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
167 |
\medskip Abstracted equivalence classes are the canonical |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
168 |
representation of elements of a quotient type. |
59031 | 169 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
170 |
|
59031 | 171 |
definition eqv_class :: "('a::partial_equiv) \<Rightarrow> 'a quot" ("\<lfloor>_\<rfloor>") |
172 |
where "\<lfloor>a\<rfloor> = Abs_quot {x. a \<sim> x}" |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
173 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
174 |
theorem quot_rep: "\<exists>a. A = \<lfloor>a\<rfloor>" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
175 |
proof (cases A) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
176 |
fix R assume R: "A = Abs_quot R" |
20811 | 177 |
assume "R \<in> quot" then have "\<exists>a. R = {x. a \<sim> x}" by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
178 |
with R have "\<exists>a. A = Abs_quot {x. a \<sim> x}" by blast |
20811 | 179 |
then show ?thesis by (unfold eqv_class_def) |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
180 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
181 |
|
20811 | 182 |
lemma quot_cases [cases type: quot]: |
183 |
obtains (rep) a where "A = \<lfloor>a\<rfloor>" |
|
184 |
using quot_rep by blast |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
185 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
186 |
|
59031 | 187 |
subsection \<open>Equality on quotients\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
188 |
|
59031 | 189 |
text \<open> |
12338
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
190 |
Equality of canonical quotient elements corresponds to the original |
de0f4a63baa5
renamed class "term" to "type" (actually "HOL.type");
wenzelm
parents:
10352
diff
changeset
|
191 |
relation as follows. |
59031 | 192 |
\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
193 |
|
59031 | 194 |
theorem eqv_class_eqI [intro]: "a \<sim> b \<Longrightarrow> \<lfloor>a\<rfloor> = \<lfloor>b\<rfloor>" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
195 |
proof - |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
196 |
assume ab: "a \<sim> b" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
197 |
have "{x. a \<sim> x} = {x. b \<sim> x}" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
198 |
proof (rule Collect_cong) |
59031 | 199 |
fix x show "a \<sim> x \<longleftrightarrow> b \<sim> x" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
200 |
proof |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
201 |
from ab have "b \<sim> a" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
202 |
also assume "a \<sim> x" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
203 |
finally show "b \<sim> x" . |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
204 |
next |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
205 |
note ab |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
206 |
also assume "b \<sim> x" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
207 |
finally show "a \<sim> x" . |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
208 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
209 |
qed |
20811 | 210 |
then show ?thesis by (simp only: eqv_class_def) |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
211 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
212 |
|
59031 | 213 |
theorem eqv_class_eqD' [dest?]: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor> \<Longrightarrow> a \<in> domain \<Longrightarrow> a \<sim> b" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
214 |
proof (unfold eqv_class_def) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
215 |
assume "Abs_quot {x. a \<sim> x} = Abs_quot {x. b \<sim> x}" |
20811 | 216 |
then have "{x. a \<sim> x} = {x. b \<sim> x}" by (simp only: Abs_quot_inject quotI) |
217 |
moreover assume "a \<in> domain" then have "a \<sim> a" .. |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
218 |
ultimately have "a \<in> {x. b \<sim> x}" by blast |
20811 | 219 |
then have "b \<sim> a" by blast |
220 |
then show "a \<sim> b" .. |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
221 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
222 |
|
59031 | 223 |
theorem eqv_class_eqD [dest?]: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor> \<Longrightarrow> a \<sim> (b::'a::equiv)" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
224 |
proof (rule eqv_class_eqD') |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
225 |
show "a \<in> domain" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
226 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
227 |
|
59031 | 228 |
lemma eqv_class_eq' [simp]: "a \<in> domain \<Longrightarrow> \<lfloor>a\<rfloor> = \<lfloor>b\<rfloor> \<longleftrightarrow> a \<sim> b" |
28616
ac1da69fbc5a
avoid accidental dependency of automated proof on sort equiv;
wenzelm
parents:
23373
diff
changeset
|
229 |
using eqv_class_eqI eqv_class_eqD' by (blast del: eqv_refl) |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
230 |
|
59031 | 231 |
lemma eqv_class_eq [simp]: "\<lfloor>a\<rfloor> = \<lfloor>b\<rfloor> \<longleftrightarrow> a \<sim> (b::'a::equiv)" |
20811 | 232 |
using eqv_class_eqI eqv_class_eqD by blast |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
233 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
234 |
|
59031 | 235 |
subsection \<open>Picking representing elements\<close> |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
236 |
|
59031 | 237 |
definition pick :: "'a::partial_equiv quot \<Rightarrow> 'a" |
238 |
where "pick A = (SOME a. A = \<lfloor>a\<rfloor>)" |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
239 |
|
59031 | 240 |
theorem pick_eqv' [intro?, simp]: "a \<in> domain \<Longrightarrow> pick \<lfloor>a\<rfloor> \<sim> a" |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
241 |
proof (unfold pick_def) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
242 |
assume a: "a \<in> domain" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
243 |
show "(SOME x. \<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>) \<sim> a" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
244 |
proof (rule someI2) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
245 |
show "\<lfloor>a\<rfloor> = \<lfloor>a\<rfloor>" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
246 |
fix x assume "\<lfloor>a\<rfloor> = \<lfloor>x\<rfloor>" |
23373 | 247 |
from this and a have "a \<sim> x" .. |
20811 | 248 |
then show "x \<sim> a" .. |
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
249 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
250 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
251 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
252 |
theorem pick_eqv [intro, simp]: "pick \<lfloor>a\<rfloor> \<sim> (a::'a::equiv)" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
253 |
proof (rule pick_eqv') |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
254 |
show "a \<in> domain" .. |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
255 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
256 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
257 |
theorem pick_inverse: "\<lfloor>pick A\<rfloor> = (A::'a::equiv quot)" |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
258 |
proof (cases A) |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
259 |
fix a assume a: "A = \<lfloor>a\<rfloor>" |
20811 | 260 |
then have "pick A \<sim> a" by simp |
261 |
then have "\<lfloor>pick A\<rfloor> = \<lfloor>a\<rfloor>" by simp |
|
10352
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
262 |
with a show ?thesis by simp |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
263 |
qed |
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
264 |
|
638e1fc6ca74
Partial equivalence relations (leftover from HOL/Quot);
wenzelm
parents:
diff
changeset
|
265 |
end |