2208
+ − 1
(* How to change the notation for \<lbrakk> \<rbrakk> meta-level implications? *)
2195
+ − 2
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 3
(*<*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 4
theory Paper
2183
+ − 5
imports "Quotient"
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 6
"LaTeXsugar"
2186
+ − 7
"../Nominal/FSet"
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 8
begin
1994
+ − 9
+ − 10
notation (latex output)
2217
+ − 11
rel_conj ("_ OOO _" [53, 53] 52) and
+ − 12
"op -->" (infix "\<rightarrow>" 100) and
+ − 13
"==>" (infix "\<Rightarrow>" 100) and
2227
+ − 14
fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
+ − 15
fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
2217
+ − 16
list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
2223
+ − 17
fempty ("\<emptyset>") and
+ − 18
funion ("_ \<union> _") and
+ − 19
finsert ("{_} \<union> _") and
2221
+ − 20
Cons ("_::_") and
+ − 21
concat ("flat") and
2223
+ − 22
fconcat ("\<Union>")
2220
+ − 23
2217
+ − 24
1994
+ − 25
2182
+ − 26
ML {*
+ − 27
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
+ − 28
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
+ − 29
let
+ − 30
val concl =
+ − 31
Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
+ − 32
in
+ − 33
case concl of (_ $ l $ r) => proj (l, r)
+ − 34
| _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
+ − 35
end);
+ − 36
*}
+ − 37
setup {*
+ − 38
Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
+ − 39
Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
+ − 40
Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
+ − 41
*}
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 42
(*>*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 43
2227
+ − 44
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 45
section {* Introduction *}
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 46
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 47
text {*
2205
+ − 48
\begin{flushright}
+ − 49
{\em ``Not using a [quotient] package has its advantages: we do not have to\\
+ − 50
collect all the theorems we shall ever want into one giant list;''}\\
2213
+ − 51
Larry Paulson \cite{Paulson06}
2237
+ − 52
\end{flushright}
2103
+ − 53
+ − 54
\noindent
2220
+ − 55
Isabelle is a popular generic theorem prover in which many logics can be
2214
+ − 56
implemented. The most widely used one, however, is Higher-Order Logic
+ − 57
(HOL). This logic consists of a small number of axioms and inference rules
+ − 58
over a simply-typed term-language. Safe reasoning in HOL is ensured by two
+ − 59
very restricted mechanisms for extending the logic: one is the definition of
+ − 60
new constants in terms of existing ones; the other is the introduction of
+ − 61
new types by identifying non-empty subsets in existing types. It is well
2223
+ − 62
understood how to use both mechanisms for dealing with quotient
2220
+ − 63
constructions in HOL (see \cite{Homeier05,Paulson06}). For example the
+ − 64
integers in Isabelle/HOL are constructed by a quotient construction over the
+ − 65
type @{typ "nat \<times> nat"} and the equivalence relation
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 66
2237
+ − 67
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 68
@{text "(n\<^isub>1, n\<^isub>2) \<approx> (m\<^isub>1, m\<^isub>2) \<equiv> n\<^isub>1 + m\<^isub>2 = m\<^isub>1 + n\<^isub>2"}\hfill\numbered{natpairequiv}
+ − 69
\end{isabelle}
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 70
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 71
\noindent
2217
+ − 72
This constructions yields the new type @{typ int} and definitions for @{text
2220
+ − 73
"0"} and @{text "1"} of type @{typ int} can be given in terms of pairs of
+ − 74
natural numbers (namely @{text "(0, 0)"} and @{text "(1, 0)"}). Operations
+ − 75
such as @{text "add"} with type @{typ "int \<Rightarrow> int \<Rightarrow> int"} can be defined in
+ − 76
terms of operations on pairs of natural numbers (namely @{text
2224
+ − 77
"add_pair (n\<^isub>1, m\<^isub>1) (n\<^isub>2,
2222
+ − 78
m\<^isub>2) \<equiv> (n\<^isub>1 + n\<^isub>2, m\<^isub>1 + m\<^isub>2)"}).
+ − 79
Similarly one can construct the type of finite sets, written @{term "\<alpha> fset"},
2223
+ − 80
by quotienting the type @{text "\<alpha> list"} according to the equivalence relation
2103
+ − 81
2237
+ − 82
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 83
@{text "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}\hfill\numbered{listequiv}
+ − 84
\end{isabelle}
2103
+ − 85
+ − 86
\noindent
2224
+ − 87
which states that two lists are equivalent if every element in one list is
+ − 88
also member in the other. The empty finite set, written @{term "{||}"}, can
+ − 89
then be defined as the empty list and the union of two finite sets, written
+ − 90
@{text "\<union>"}, as list append.
2220
+ − 91
+ − 92
An area where quotients are ubiquitous is reasoning about programming language
2226
+ − 93
calculi. A simple example is the lambda-calculus, whose raw terms are defined as
2220
+ − 94
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 95
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 96
@{text "t ::= x | t t | \<lambda>x.t"}\hfill\numbered{lambda}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 97
\end{isabelle}
2217
+ − 98
2220
+ − 99
\noindent
2222
+ − 100
The problem with this definition arises when one attempts to
+ − 101
prove formally, for example, the substitution lemma \cite{Barendregt81} by induction
+ − 102
over the structure of terms. This can be fiendishly complicated (see
2220
+ − 103
\cite[Pages 94--104]{CurryFeys58} for some ``rough'' sketches of a proof
2226
+ − 104
about raw lambda-terms). In contrast, if we reason about
2220
+ − 105
$\alpha$-equated lambda-terms, that means terms quotient according to
2223
+ − 106
$\alpha$-equivalence, then the reasoning infrastructure provided,
+ − 107
for example, by Nominal Isabelle \cite{UrbanKaliszyk11} makes the formal
2222
+ − 108
proof of the substitution lemma almost trivial.
2220
+ − 109
+ − 110
The difficulty is that in order to be able to reason about integers, finite
2221
+ − 111
sets or $\alpha$-equated lambda-terms one needs to establish a reasoning
2220
+ − 112
infrastructure by transferring, or \emph{lifting}, definitions and theorems
2226
+ − 113
from the raw type @{typ "nat \<times> nat"} to the quotient type @{typ int}
2220
+ − 114
(similarly for finite sets and $\alpha$-equated lambda-terms). This lifting
2222
+ − 115
usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.
2223
+ − 116
It is feasible to to this work manually, if one has only a few quotient
+ − 117
constructions at hand. But if they have to be done over and over again as in
2222
+ − 118
Nominal Isabelle, then manual reasoning is not an option.
2221
+ − 119
2223
+ − 120
The purpose of a \emph{quotient package} is to ease the lifting of theorems
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 121
and automate the reasoning as much as possible. In the
2223
+ − 122
context of HOL, there have been a few quotient packages already
2234
+ − 123
\cite{harrison-thesis,Slotosch97}. The most notable one is by Homeier
2223
+ − 124
\cite{Homeier05} implemented in HOL4. The fundamental construction these
+ − 125
quotient packages perform can be illustrated by the following picture:
2217
+ − 126
+ − 127
\begin{center}
2220
+ − 128
\mbox{}\hspace{20mm}\begin{tikzpicture}
+ − 129
%%\draw[step=2mm] (-4,-1) grid (4,1);
+ − 130
+ − 131
\draw[very thick] (0.7,0.3) circle (4.85mm);
+ − 132
\draw[rounded corners=1mm, very thick] ( 0.0,-0.9) rectangle ( 1.8, 0.9);
+ − 133
\draw[rounded corners=1mm, very thick] (-1.95,0.8) rectangle (-2.9,-0.195);
+ − 134
+ − 135
\draw (-2.0, 0.8) -- (0.7,0.8);
+ − 136
\draw (-2.0,-0.195) -- (0.7,-0.195);
2103
+ − 137
2220
+ − 138
\draw ( 0.7, 0.23) node {\begin{tabular}{@ {}c@ {}}equiv-\\[-1mm]clas.\end{tabular}};
+ − 139
\draw (-2.45, 0.35) node {\begin{tabular}{@ {}c@ {}}new\\[-1mm]type\end{tabular}};
+ − 140
\draw (1.8, 0.35) node[right=-0.1mm]
2224
+ − 141
{\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw elements)\end{tabular}};
2220
+ − 142
\draw (0.9, -0.55) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
+ − 143
+ − 144
\draw[->, very thick] (-1.8, 0.36) -- (-0.1,0.36);
+ − 145
\draw[<-, very thick] (-1.8, 0.16) -- (-0.1,0.16);
2221
+ − 146
\draw (-0.95, 0.26) node[above=0.4mm] {@{text Rep}};
+ − 147
\draw (-0.95, 0.26) node[below=0.4mm] {@{text Abs}};
2220
+ − 148
+ − 149
\end{tikzpicture}
+ − 150
\end{center}
2217
+ − 151
2221
+ − 152
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 153
The starting point is an existing type, to which we often refer as the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 154
\emph{raw type}, over which an equivalence relation given by the user is
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 155
defined. With this input the package introduces a new type, to which we
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
refer as the \emph{quotient type}. This type comes with an
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 157
\emph{abstraction} and a \emph{representation} function, written @{text Abs}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 158
and @{text Rep}.\footnote{Actually slightly more basic functions are given;
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 159
the functions @{text Abs} and @{text Rep} need to be derived from them. We
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 160
will show the details later. } These functions relate elements in the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 161
existing type to ones in the new type and vice versa; they can be uniquely
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 162
identified by their type. For example for the integer quotient construction
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 163
the types of @{text Abs} and @{text Rep} are
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 164
2224
+ − 165
+ − 166
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2234
+ − 167
@{text "Abs :: nat \<times> nat \<Rightarrow> int"}\hspace{10mm}@{text "Rep :: int \<Rightarrow> nat \<times> nat"}
2224
+ − 168
\end{isabelle}
+ − 169
+ − 170
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 171
We therefore often write @{text Abs_int} and @{text Rep_int} if the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 172
typing information is important.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 173
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 174
Every abstraction and representation function stands for an isomorphism
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 175
between the non-empty subset and elements in the new type. They are
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 176
necessary for making definitions involving the new type. For example @{text
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 177
"0"} and @{text "1"} of type @{typ int} can be defined as
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 178
2221
+ − 179
+ − 180
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 181
@{text "0 \<equiv> Abs_int (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs_int (1, 0)"}
2221
+ − 182
\end{isabelle}
+ − 183
+ − 184
\noindent
2224
+ − 185
Slightly more complicated is the definition of @{text "add"} having type
2222
+ − 186
@{typ "int \<Rightarrow> int \<Rightarrow> int"}. Its definition is as follows
2221
+ − 187
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 188
@{text [display, indent=10] "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
2221
+ − 189
+ − 190
\noindent
2224
+ − 191
where we take the representation of the arguments @{text n} and @{text m},
2226
+ − 192
add them according to the function @{text "add_pair"} and then take the
2221
+ − 193
abstraction of the result. This is all straightforward and the existing
+ − 194
quotient packages can deal with such definitions. But what is surprising
2223
+ − 195
that none of them can deal with slightly more complicated definitions involving
2221
+ − 196
\emph{compositions} of quotients. Such compositions are needed for example
2226
+ − 197
in case of quotienting lists to obtain finite sets and the operator that
+ − 198
flattens lists of lists, defined as follows
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 199
2220
+ − 200
@{thm [display, indent=10] concat.simps(1) concat.simps(2)[no_vars]}
2183
+ − 201
2103
+ − 202
\noindent
2221
+ − 203
We expect that the corresponding operator on finite sets, written @{term "fconcat"},
2234
+ − 204
builds the union of finite sets of finite sets:
2188
+ − 205
2220
+ − 206
@{thm [display, indent=10] fconcat_empty[no_vars] fconcat_insert[no_vars]}
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 207
2103
+ − 208
\noindent
2223
+ − 209
The quotient package should provide us with a definition for @{text "\<Union>"} in
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 210
terms of @{text flat}, @{text Rep_fset} and @{text Abs_fset}. The problem is
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 211
that the method used in the existing quotient
2224
+ − 212
packages of just taking the representation of the arguments and then take
+ − 213
the abstraction of the result is \emph{not} enough. The reason is that case in case
+ − 214
of @{text "\<Union>"} we obtain the incorrect definition
+ − 215
2234
+ − 216
@{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat (Rep_fset S))"}
2221
+ − 217
2224
+ − 218
\noindent
+ − 219
where the right-hand side is not even typable! This problem can be remedied in the
+ − 220
existing quotient packages by introducing an intermediate step and reasoning
2226
+ − 221
about flattening of lists of finite sets. However, this remedy is rather
2224
+ − 222
cumbersome and inelegant in light of our work, which can deal with such
+ − 223
definitions directly. The solution is that we need to build aggregate
+ − 224
representation and abstraction functions, which in case of @{text "\<Union>"}
+ − 225
generate the following definition
+ − 226
2234
+ − 227
@{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat ((map Rep_fset \<circ> Rep_fset) S))"}
2221
+ − 228
+ − 229
\noindent
2223
+ − 230
where @{term map} is the usual mapping function for lists. In this paper we
2224
+ − 231
will present a formal definition of our aggregate abstraction and
2223
+ − 232
representation functions (this definition was omitted in \cite{Homeier05}).
2224
+ − 233
They generate definitions, like the one above for @{text "\<Union>"},
2226
+ − 234
according to the type of the raw constant and the type
2224
+ − 235
of the quotient constant. This means we also have to extend the notions
2237
+ − 236
of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
2231
+ − 237
from Homeier \cite{Homeier05}.
2223
+ − 238
2237
+ − 239
We are also able to address the criticism by Paulson \cite{Paulson06} cited
+ − 240
at the beginning of this section about having to collect theorems that are
+ − 241
lifted from the raw level to the quotient level. Our quotient package is the
+ − 242
first one that is modular so that it allows to lift single theorems
+ − 243
separately. This has the advantage for the user to develop a formal theory
+ − 244
interactively an a natural progression. A pleasing result of the modularity
+ − 245
is also that we are able to clearly specify what needs to be done in the
+ − 246
lifting process (this was only hinted at in \cite{Homeier05} and implemented
+ − 247
as a ``rough recipe'' in ML-code).
+ − 248
+ − 249
The paper is organised as follows: Section \ref{sec:prelims} presents briefly
+ − 250
some necessary preliminaries; Section \ref{sec:type} presents the definitions
+ − 251
of quotient types and shows how definitions can be made over quotient types. \ldots
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 252
*}
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 253
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 254
2224
+ − 255
section {* Preliminaries and General Quotient\label{sec:prelims} *}
1978
+ − 256
+ − 257
text {*
2182
+ − 258
In this section we present the definitions of a quotient that follow
+ − 259
those by Homeier, the proofs can be found there.
+ − 260
+ − 261
\begin{definition}[Quotient]
+ − 262
A relation $R$ with an abstraction function $Abs$
+ − 263
and a representation function $Rep$ is a \emph{quotient}
+ − 264
if and only if:
1978
+ − 265
2182
+ − 266
\begin{enumerate}
+ − 267
\item @{thm (rhs1) Quotient_def[of "R", no_vars]}
+ − 268
\item @{thm (rhs2) Quotient_def[of "R", no_vars]}
+ − 269
\item @{thm (rhs3) Quotient_def[of "R", no_vars]}
+ − 270
\end{enumerate}
+ − 271
+ − 272
\end{definition}
1978
+ − 273
2188
+ − 274
\begin{definition}[Relation map and function map]\\
2190
+ − 275
@{thm fun_rel_def[of "R1" "R2", no_vars]}\\
2182
+ − 276
@{thm fun_map_def[no_vars]}
+ − 277
\end{definition}
+ − 278
+ − 279
The main theorems for building higher order quotients is:
+ − 280
\begin{lemma}[Function Quotient]
+ − 281
If @{thm (prem 1) fun_quotient[no_vars]} and @{thm (prem 2) fun_quotient[no_vars]}
+ − 282
then @{thm (concl) fun_quotient[no_vars]}
+ − 283
\end{lemma}
+ − 284
1978
+ − 285
*}
+ − 286
2195
+ − 287
subsection {* Higher Order Logic *}
+ − 288
+ − 289
text {*
+ − 290
+ − 291
Types:
+ − 292
\begin{eqnarray}\nonumber
+ − 293
@{text "\<sigma> ::="} & @{text "\<alpha>"} & \textrm{(type variable)} \\ \nonumber
+ − 294
@{text "|"} & @{text "(\<sigma>,\<dots>,\<sigma>)\<kappa>"} & \textrm{(type construction)}
+ − 295
\end{eqnarray}
+ − 296
+ − 297
Terms:
+ − 298
\begin{eqnarray}\nonumber
+ − 299
@{text "t ::="} & @{text "x\<^isup>\<sigma>"} & \textrm{(variable)} \\ \nonumber
+ − 300
@{text "|"} & @{text "c\<^isup>\<sigma>"} & \textrm{(constant)} \\ \nonumber
+ − 301
@{text "|"} & @{text "t t"} & \textrm{(application)} \\ \nonumber
+ − 302
@{text "|"} & @{text "\<lambda>x\<^isup>\<sigma>. t"} & \textrm{(abstraction)} \\ \nonumber
+ − 303
\end{eqnarray}
+ − 304
2234
+ − 305
{\it Say more about containers / maping functions }
+ − 306
2237
+ − 307
Such maps for most common types (list, pair, sum,
+ − 308
option, \ldots) are described in Homeier, and we assume that @{text "map"}
+ − 309
is the function that returns a map for a given type.
+ − 310
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 311
{\it say something about our use of @{text "\<sigma>s"}}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 312
2195
+ − 313
*}
+ − 314
2237
+ − 315
section {* Quotient Types and Quotient Definitions\label{sec:type} *}
1978
+ − 316
2234
+ − 317
text {*
+ − 318
The first step in a quotient constructions is to take a name for the new
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 319
type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 320
defined over a raw type, say @{text "\<sigma>"}. The type of this equivalence
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 321
relation must be of type @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 322
the declaration is therefore
2234
+ − 323
2235
+ − 324
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 325
\isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}
+ − 326
\end{isabelle}
+ − 327
+ − 328
\noindent
2237
+ − 329
and a proof that @{text "R"} is indeed an equivalence relation. Two concrete
+ − 330
examples are
+ − 331
+ − 332
+ − 333
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 334
\begin{tabular}{@ {}l}
+ − 335
\isacommand{quotient\_type}~~@{text "int = nat \<times> nat / \<approx>\<^bsub>nat \<times> nat\<^esub>"}\\
+ − 336
\isacommand{quotient\_type}~~@{text "\<alpha> fset = \<alpha> list / \<approx>\<^bsub>list\<^esub>"}
+ − 337
\end{tabular}
+ − 338
\end{isabelle}
+ − 339
+ − 340
\noindent
+ − 341
which introduce the type of integers and of finite sets using the
+ − 342
equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
+ − 343
"\<approx>\<^bsub>list\<^esub>"} defined earlier in \eqref{natpairequiv} and
+ − 344
\eqref{listequiv}, respectively. Given this data, we declare internally
+ − 345
the quotient types as
2234
+ − 346
+ − 347
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 348
\isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
+ − 349
\end{isabelle}
+ − 350
+ − 351
\noindent
2237
+ − 352
where the right hand side is the (non-empty) set of equivalence classes of
+ − 353
@{text "R"}. The restriction in this declaration is that the type variables
+ − 354
in the raw type @{text "\<sigma>"} must be included in the type variables @{text
+ − 355
"\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will provide us with
+ − 356
abstraction and representation functions having the type
2182
+ − 357
2234
+ − 358
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 359
@{text "abs_\<kappa>\<^isub>q :: \<sigma> set \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}\hspace{10mm}@{text "rep_\<kappa>\<^isub>q :: \<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma> set"}
+ − 360
\end{isabelle}
+ − 361
2235
+ − 362
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 363
They relate the new quotient type and equivalence classes of the raw
2235
+ − 364
type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
+ − 365
to work with the following derived abstraction and representation functions
+ − 366
2234
+ − 367
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 368
@{text "Abs_\<kappa>\<^isub>q x \<equiv> abs_\<kappa>\<^isub>q (R x)"}\hspace{10mm}@{text "Rep_\<kappa>\<^isub>q x \<equiv> \<epsilon> (rep_\<kappa>\<^isub>q x)"}
+ − 369
\end{isabelle}
+ − 370
+ − 371
\noindent
2235
+ − 372
on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"} in the
2237
+ − 373
definition of @{text "Rep_\<kappa>\<^isub>q"}. These derived notions relate the
+ − 374
quotient type and the raw type directly, as can be seen from their type,
+ − 375
namely @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"},
+ − 376
respectively. Given that @{text "R"} is an equivalence relation, the
+ − 377
following property
+ − 378
2234
+ − 379
+ − 380
@{text [display, indent=10] "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}
+ − 381
+ − 382
\noindent
+ − 383
holds (for the proof see \cite{Homeier05}).
2182
+ − 384
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 385
The next step in a quotient construction is to introduce new definitions
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 386
involving the quotient type, which need to be defined in terms of concepts
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 387
of the raw type (remember this is the only way how to extend HOL
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 388
with new definitions). For the user visible is the declaration
2235
+ − 389
+ − 390
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2237
+ − 391
\isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
2235
+ − 392
\end{isabelle}
+ − 393
2237
+ − 394
\noindent
+ − 395
where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
+ − 396
and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
+ − 397
given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 398
in places where a quotient and raw type are involved). Two concrete examples are
2188
+ − 399
2237
+ − 400
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
+ − 401
\begin{tabular}{@ {}l}
+ − 402
\isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
+ − 403
\isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
+ − 404
\isacommand{is}~~@{text "flat"}
+ − 405
\end{tabular}
+ − 406
\end{isabelle}
+ − 407
+ − 408
\noindent
+ − 409
The first one declares zero for integers and the second the operator for
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 410
building unions of finite sets.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 411
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 412
The problem for us is that from such declarations we need to derive proper
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 413
definitions using the @{text "Abs"} and @{text "Rep"} functions for the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 414
quotient types involved. The data we rely on is the given quotient type
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 415
@{text "\<tau>"} and the raw type @{text "\<sigma>"}. They allow us to define aggregate
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 416
abstraction and representation functions using the functions @{text "ABS (\<sigma>,
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 417
\<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses are given below. The idea behind
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 418
them is to recursively descend into types @{text \<sigma>} and @{text \<tau>}, and generate the appropriate
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 419
@{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
2244
+ − 420
we return just the identity whenever the types are equal. All clauses
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 421
are as follows:
2182
+ − 422
2227
+ − 423
\begin{center}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 424
\begin{tabular}{rcl}
2227
+ − 425
\multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 426
@{text "ABS (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\\
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 427
@{text "REP (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\smallskip\\
2227
+ − 428
\multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\
2233
+ − 429
@{text "ABS (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "REP (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> ABS (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\\
+ − 430
@{text "REP (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> REP (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\smallskip\\
2227
+ − 431
\multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\
2232
+ − 432
@{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (ABS (\<sigma>s, \<tau>s))"}\\
+ − 433
@{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (REP (\<sigma>s, \<tau>s))"}\smallskip\\
2227
+ − 434
\multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors:}\\
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 435
@{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "Abs_\<kappa>\<^isub>q \<circ> (MAP(\<rho>s \<kappa>) (ABS (\<sigma>s', \<tau>s)))"}\\
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 436
@{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "(MAP(\<rho>s \<kappa>) (REP (\<sigma>s', \<tau>s))) \<circ> Rep_\<kappa>\<^isub>q"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 437
\end{tabular}
2227
+ − 438
\end{center}
2234
+ − 439
%
2232
+ − 440
\noindent
2237
+ − 441
where in the last two clauses we have that the quotient type @{text "\<alpha>s
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 442
\<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
2237
+ − 443
@{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
+ − 444
list"}). The quotient construction ensures that the type variables in @{text
+ − 445
"\<rho>s"} must be amongst the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 446
matchers for the @{text "\<alpha>s"} when matching @{text "\<rho>s \<kappa>"} against
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 447
@{text "\<sigma>s \<kappa>"}. The
2237
+ − 448
function @{text "MAP"} calculates an \emph{aggregate map-function} for a raw
+ − 449
type as follows:
+ − 450
%
2227
+ − 451
\begin{center}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 452
\begin{tabular}{rcl}
2237
+ − 453
@{text "MAP' (\<alpha>)"} & $\dn$ & @{text "a\<^sup>\<alpha>"}\\
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 454
@{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id :: \<kappa> \<Rightarrow> \<kappa>"}\\
2232
+ − 455
@{text "MAP' (\<sigma>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (MAP'(\<sigma>s))"}\smallskip\\
2233
+ − 456
@{text "MAP (\<sigma>)"} & $\dn$ & @{text "\<lambda>as. MAP'(\<sigma>)"}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 457
\end{tabular}
2227
+ − 458
\end{center}
2237
+ − 459
%
2232
+ − 460
\noindent
2233
+ − 461
In this definition we abuse the fact that we can interpret type-variables @{text \<alpha>} as
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 462
term variables @{text a}. In the last clause we build an abstraction over all
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 463
term-variables inside the aggregate map-function generated by the auxiliary function
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 464
@{text "MAP'"}.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 465
The need of aggregate map-functions can be appreciated if we build quotients,
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 466
say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types of the form @{text "(\<alpha> list) \<times> \<beta>"}.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 467
In this case @{text MAP} generates the aggregate map-function:
2232
+ − 468
2233
+ − 469
@{text [display, indent=10] "\<lambda>a b. map_prod (map a) b"}
+ − 470
+ − 471
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 472
which we need to define the aggregate abstraction and representation functions.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 473
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 474
To se how these definitions pan out in practise, let us return to our
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 475
example about @{term "concat"} and @{term "fconcat"}, where we have types
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 476
@{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 477
fset"}. Feeding them into @{text ABS} gives us the abstraction function
2233
+ − 478
+ − 479
@{text [display, indent=10] "(map (map id \<circ> Rep_fset) \<circ> Rep_fset) \<singlearr> Abs_fset \<circ> map id"}
+ − 480
+ − 481
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 482
after some @{text "\<beta>"}-simplifications. In our implementation we further
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 483
simplify this abstraction function employing the usual laws about @{text
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 484
"map"}s and @{text "id"}, namely @{term "map id = id"} and @{text "f \<circ> id =
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 485
id \<circ> f = f"}. This gives us the abstraction function
2237
+ − 486
2233
+ − 487
@{text [display, indent=10] "(map Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset"}
+ − 488
+ − 489
\noindent
+ − 490
which we can use for defining @{term "fconcat"} as follows
+ − 491
+ − 492
@{text [display, indent=10] "\<Union> \<equiv> ((map Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset) flat"}
2232
+ − 493
2237
+ − 494
\noindent
+ − 495
Note that by using the operator @{text "\<singlearr>"} we do not have to
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 496
distinguish between arguments and results: the representation and abstraction
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 497
functions are just inverses of each other, which we can combine using
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 498
@{text "\<singlearr>"} to deal uniformly with arguments of functions and
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 499
their result. As a result, all definitions in the quotient package
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 500
are of the general form
2188
+ − 501
2237
+ − 502
@{text [display, indent=10] "c \<equiv> ABS (\<sigma>, \<tau>) t"}
2227
+ − 503
2237
+ − 504
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 505
where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 506
type of the defined quotient constant @{text "c"}. To ensure we
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 507
obtained a correct definition, we can prove:
2227
+ − 508
+ − 509
\begin{lemma}
+ − 510
If @{text "ABS (\<sigma>, \<tau>)"} returns some abstraction function @{text "Abs"}
+ − 511
and @{text "REP (\<sigma>, \<tau>)"} some representation function @{text "Rep"},
+ − 512
then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
+ − 513
@{text "\<tau> \<Rightarrow> \<sigma>"}.
+ − 514
\end{lemma}
2233
+ − 515
2237
+ − 516
\begin{proof}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 517
By induction and analysing the definitions of @{text "ABS"}, @{text "REP"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 518
and @{text "MAP"}. The cases of equal and function types are straightforward
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 519
(the latter follows from @{text "\<singlearr>"} having the type
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 520
@{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). The case of equal
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 521
type constructors follows by observing that a map-function after applying
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 522
the functions @{text "ABS (\<sigma>s, \<tau>s)"} produce a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 523
The interesting case is the one with unequal type constructors. Since we know the
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 524
quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have that @{text "Abs_\<kappa>\<^isub>q"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 525
is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}, that can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 526
where the type variables @{text "\<alpha>s"} are instantiated with @{text "\<tau>s"}. The
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 527
complete type can be calculated by observing that @{text "MAP (\<rho>s \<kappa>)"} after applying
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 528
the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it, returns a term of type
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 529
@{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is equivalent to @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 530
as desired.\qed
2237
+ − 531
\end{proof}
+ − 532
+ − 533
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 534
The reader should note that this lemma fails for the abstraction and representation
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 535
functions used, for example, in Homeier's quotient package.
2188
+ − 536
*}
+ − 537
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 538
section {* Respectfulness and Preservation *}
2188
+ − 539
+ − 540
text {*
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 541
Before we can lift theorems involving the raw types to theorems over
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 542
quotient types, we have to impose some restrictions. The reason is that not
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 543
all theorems can be lifted. Homeier formulates these restrictions in terms
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 544
of \emph{respectfullness} and \emph{preservation} of constants occuring in
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 545
theorems.
2188
+ − 546
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 547
The respectfulness property for a constant states that it essentially
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 548
respects the equivalence relation involved in the quotient. An example
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 549
is the function returning bound variables of a lambda-term (see \eqref{lambda})
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 550
and @{text "\<alpha>"}-equivalence. It will turn out that this function is not
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 551
respectful.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 552
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 553
To state the respectfulness property we have to define \emph{aggregate equivalence
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 554
relations}.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 555
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 556
@{text [display] "GIVE DEFINITION HERE"}
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 557
2188
+ − 558
class returned by this constant depends only on the equivalence
2207
+ − 559
classes of the arguments applied to the constant. To automatically
+ − 560
lift a theorem that talks about a raw constant, to a theorem about
+ − 561
the quotient type a respectfulness theorem is required.
+ − 562
+ − 563
A respectfulness condition for a constant can be expressed in
+ − 564
terms of an aggregate relation between the constant and itself,
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 565
for example the respectfullness for @{text "append"}
2188
+ − 566
can be stated as:
+ − 567
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 568
@{text [display, indent=10] "(R \<doublearr> R \<doublearr> R) append append"}
2182
+ − 569
2190
+ − 570
\noindent
2228
+ − 571
Which after unfolding the definition of @{term "op ===>"} is equivalent to:
2188
+ − 572
2228
+ − 573
@{thm [display, indent=10] append_rsp_unfolded[no_vars]}
2188
+ − 574
2228
+ − 575
\noindent An aggregate relation is defined in terms of relation
+ − 576
composition, so we define it first:
2188
+ − 577
+ − 578
\begin{definition}[Composition of Relations]
2190
+ − 579
@{abbrev "rel_conj R1 R2"} where @{text OO} is the predicate
+ − 580
composition @{thm pred_compI[no_vars]}
2188
+ − 581
\end{definition}
+ − 582
2207
+ − 583
The aggregate relation for an aggregate raw type and quotient type
+ − 584
is defined as:
2188
+ − 585
+ − 586
\begin{itemize}
2207
+ − 587
\item @{text "REL(\<alpha>\<^isub>1, \<alpha>\<^isub>2)"} = @{text "op ="}
+ − 588
\item @{text "REL(\<sigma>, \<sigma>)"} = @{text "op ="}
+ − 589
\item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>n))\<kappa>)"} = @{text "(rel \<kappa>) (REL(\<sigma>\<^isub>1,\<tau>\<^isub>1)) \<dots> (REL(\<sigma>\<^isub>n,\<tau>\<^isub>n))"}
+ − 590
\item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>\<^isub>1, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>m))\<kappa>\<^isub>2)"} = @{text "(rel \<kappa>\<^isub>1) (REL(\<rho>\<^isub>1,\<nu>\<^isub>1) \<dots> (REL(\<rho>\<^isub>p,\<nu>\<^isub>p) OOO Eqv_\<kappa>\<^isub>2"} provided @{text "\<eta> \<kappa>\<^isub>2 = (\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>p)\<kappa>\<^isub>1 \<and> \<exists>s. s(\<sigma>s\<kappa>\<^isub>1)=\<rho>s\<kappa>\<^isub>1 \<and> s(\<tau>s\<kappa>\<^isub>2)=\<nu>s\<kappa>\<^isub>2"}
2189
+ − 591
2188
+ − 592
\end{itemize}
+ − 593
2207
+ − 594
Again, the last case is novel, so lets look at the example of
+ − 595
respectfullness for @{term concat}. The statement according to
+ − 596
the definition above is:
2190
+ − 597
2228
+ − 598
@{thm [display, indent=10] concat_rsp[no_vars]}
2189
+ − 599
2190
+ − 600
\noindent
+ − 601
By unfolding the definition of relation composition and relation map
+ − 602
we can see the equivalent statement just using the primitive list
+ − 603
equivalence relation:
+ − 604
2228
+ − 605
@{thm [display, indent=10] concat_rsp_unfolded[of "a" "a'" "b'" "b", no_vars]}
2189
+ − 606
2190
+ − 607
The statement reads that, for any lists of lists @{term a} and @{term b}
+ − 608
if there exist intermediate lists of lists @{term "a'"} and @{term "b'"}
+ − 609
such that each element of @{term a} is in the relation with an appropriate
+ − 610
element of @{term a'}, @{term a'} is in relation with @{term b'} and each
+ − 611
element of @{term b'} is in relation with the appropriate element of
+ − 612
@{term b}.
2189
+ − 613
+ − 614
*}
+ − 615
+ − 616
2190
+ − 617
text {*
2228
+ − 618
Sometimes a non-lifted polymorphic constant is instantiated to a
+ − 619
type being lifted. For example take the @{term "op #"} which inserts
+ − 620
an element in a list of pairs of natural numbers. When the theorem
+ − 621
is lifted, the pairs of natural numbers are to become integers, but
+ − 622
the head constant is still supposed to be the head constant, just
+ − 623
with a different type. To be able to lift such theorems
+ − 624
automatically, additional theorems provided by the user are
+ − 625
necessary, we call these \emph{preservation} theorems following
+ − 626
Homeier's naming.
2196
+ − 627
+ − 628
To lift theorems that talk about insertion in lists of lifted types
+ − 629
we need to know that for any quotient type with the abstraction and
+ − 630
representation functions @{text "Abs"} and @{text Rep} we have:
+ − 631
2228
+ − 632
@{thm [display, indent=10] (concl) cons_prs[no_vars]}
2196
+ − 633
+ − 634
This is not enough to lift theorems that talk about quotient compositions.
+ − 635
For some constants (for example empty list) it is possible to show a
+ − 636
general compositional theorem, but for @{term "op #"} it is necessary
+ − 637
to show that it respects the particular quotient type:
+ − 638
2228
+ − 639
@{thm [display, indent=10] insert_preserve2[no_vars]}
2190
+ − 640
*}
+ − 641
2191
+ − 642
subsection {* Composition of Quotient theorems *}
2189
+ − 643
2191
+ − 644
text {*
+ − 645
Given two quotients, one of which quotients a container, and the
+ − 646
other quotients the type in the container, we can write the
2193
+ − 647
composition of those quotients. To compose two quotient theorems
2207
+ − 648
we compose the relations with relation composition as defined above
+ − 649
and the abstraction and relation functions are the ones of the sub
+ − 650
quotients composed with the usual function composition.
+ − 651
The @{term "Rep"} and @{term "Abs"} functions that we obtain agree
+ − 652
with the definition of aggregate Abs/Rep functions and the
2193
+ − 653
relation is the same as the one given by aggregate relations.
+ − 654
This becomes especially interesting
2191
+ − 655
when we compose the quotient with itself, as there is no simple
+ − 656
intermediate step.
+ − 657
2242
+ − 658
Lets take again the example of @{term flat}. To be able to lift
2207
+ − 659
theorems that talk about it we provide the composition quotient
+ − 660
theorems, which then lets us perform the lifting procedure in an
+ − 661
unchanged way:
2188
+ − 662
2190
+ − 663
@{thm [display] quotient_compose_list[no_vars]}
2192
+ − 664
*}
+ − 665
2191
+ − 666
2227
+ − 667
section {* Lifting of Theorems *}
1978
+ − 668
2194
+ − 669
text {*
+ − 670
The core of the quotient package takes an original theorem that
+ − 671
talks about the raw types, and the statement of the theorem that
+ − 672
it is supposed to produce. This is different from other existing
2207
+ − 673
quotient packages, where only the raw theorems were necessary.
2194
+ − 674
We notice that in some cases only some occurrences of the raw
+ − 675
types need to be lifted. This is for example the case in the
+ − 676
new Nominal package, where a raw datatype that talks about
+ − 677
pairs of natural numbers or strings (being lists of characters)
+ − 678
should not be changed to a quotient datatype with constructors
+ − 679
taking integers or finite sets of characters. To simplify the
+ − 680
use of the quotient package we additionally provide an automated
+ − 681
statement translation mechanism that replaces occurrences of
+ − 682
types that match given quotients by appropriate lifted types.
+ − 683
+ − 684
Lifting the theorems is performed in three steps. In the following
+ − 685
we call these steps \emph{regularization}, \emph{injection} and
+ − 686
\emph{cleaning} following the names used in Homeier's HOL
2197
+ − 687
implementation.
2193
+ − 688
2197
+ − 689
We first define the statement of the regularized theorem based
+ − 690
on the original theorem and the goal theorem. Then we define
+ − 691
the statement of the injected theorem, based on the regularized
2208
+ − 692
theorem and the goal. We then show the 3 proofs, as all three
2197
+ − 693
can be performed independently from each other.
2193
+ − 694
2194
+ − 695
*}
1994
+ − 696
2197
+ − 697
subsection {* Regularization and Injection statements *}
1994
+ − 698
+ − 699
text {*
2197
+ − 700
2207
+ − 701
We first define the function @{text REG}, which takes the statements
+ − 702
of the raw theorem and the lifted theorem (both as terms) and
+ − 703
returns the statement of the regularized version. The intuition
+ − 704
behind this function is that it replaces quantifiers and
+ − 705
abstractions involving raw types by bounded ones, and equalities
+ − 706
involving raw types are replaced by appropriate aggregate
+ − 707
relations. It is defined as follows:
1994
+ − 708
2244
+ − 709
\begin{center}
+ − 710
\begin{tabular}{rcl}
+ − 711
\multicolumn{3}{@ {\hspace{-4mm}}l}{abstractions (with same types and different types):}\\
+ − 712
@{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<sigma>. s)"} & $\dn$ & @{text "\<lambda>x : \<sigma>. REG (t, s)"}\\
+ − 713
@{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<tau>. s)"} & $\dn$ & @{text "\<lambda>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}\\
+ − 714
\multicolumn{3}{@ {\hspace{-4mm}}l}{quantification (over same types and different types):}\\
+ − 715
@{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<sigma>. s)"} & $\dn$ & @{text "\<forall>x : \<sigma>. REG (t, s)"}\\
+ − 716
@{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<tau>. s)"} & $\dn$ & @{text "\<forall>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}\\
2245
+ − 717
\multicolumn{3}{@ {\hspace{-4mm}}l}{equalities (with same types and different types):}\\
2244
+ − 718
@{text "REG ((op =) : \<sigma>, (op =) : \<sigma>)"} & $\dn$ & @{text "(op =) : \<sigma>"}\\
+ − 719
@{text "REG ((op =) : \<sigma>, (op =) : \<tau>)"} & $\dn$ & @{text "REL (\<sigma>, \<tau>) : \<sigma>"}\\
2245
+ − 720
\multicolumn{3}{@ {\hspace{-4mm}}l}{applications, variables, constants:}\\
2244
+ − 721
@{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2)"} & $\dn$ & @{text "REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}\\
+ − 722
@{text "REG (v\<^isub>1, v\<^isub>2)"} & $\dn$ & @{text "v\<^isub>1"}\\
+ − 723
@{text "REG (c\<^isub>1, c\<^isub>2)"} & $\dn$ & @{text "c\<^isub>1"}\\
+ − 724
\end{tabular}
+ − 725
\end{center}
1994
+ − 726
2230
+ − 727
In the above definition we omitted the cases for existential quantifiers
2207
+ − 728
and unique existential quantifiers, as they are very similar to the cases
+ − 729
for the universal quantifier.
2197
+ − 730
2207
+ − 731
Next we define the function @{text INJ} which takes the statement of
+ − 732
the regularized theorems and the statement of the lifted theorem both as
2230
+ − 733
terms and returns the statement of the injected theorem:
2198
+ − 734
2245
+ − 735
\begin{center}
+ − 736
\begin{tabular}{rcl}
+ − 737
\multicolumn{3}{@ {\hspace{-4mm}}l}{abstractions (with same types and different types):}\\
+ − 738
@{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<sigma>) "} & $\dn$ & @{text "\<lambda>x. INJ (t, s)"}\\
+ − 739
@{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<tau>) "} & $\dn$ & @{text "REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x. (INJ (t, s))))"}\\
+ − 740
@{text "INJ ((\<lambda>x \<in> R. t) : \<sigma>, (\<lambda>x. s) : \<tau>) "} & $\dn$ & @{text "REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x \<in> R. (INJ (t, s))))"}\\
+ − 741
\multicolumn{3}{@ {\hspace{-4mm}}l}{quantification (over same types and different types):}\\
+ − 742
@{text "INJ (\<forall> t, \<forall> s) "} & $\dn$ & @{text "\<forall> (INJ (t, s))"}\\
+ − 743
@{text "INJ (\<forall> t \<in> R, \<forall> s) "} & $\dn$ & @{text "\<forall> INJ (t, s) \<in> R"}\\
+ − 744
\multicolumn{3}{@ {\hspace{-4mm}}l}{applications, variables, constants:}\\
+ − 745
@{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) "} & $\dn$ & @{text " INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}\\
+ − 746
@{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<sigma>) "} & $\dn$ & @{text "v\<^isub>1"}\\
+ − 747
@{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<tau>) "} & $\dn$ & @{text "REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (v\<^isub>1))"}\\
+ − 748
@{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<sigma>) "} & $\dn$ & @{text "c\<^isub>1"}\\
+ − 749
@{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<tau>) "} & $\dn$ & @{text "REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (c\<^isub>1))"}\\
+ − 750
\end{tabular}
+ − 751
\end{center}
2198
+ − 752
+ − 753
For existential quantifiers and unique existential quantifiers it is
2230
+ − 754
defined similarly to the universal one.
2198
+ − 755
2197
+ − 756
*}
+ − 757
2208
+ − 758
subsection {* Proof procedure *}
+ − 759
2242
+ − 760
(* In the below the type-guiding 'QuotTrue' assumption is removed. We need it
+ − 761
only for bound variables without types, while in the paper presentation
+ − 762
variables are typed *)
2197
+ − 763
+ − 764
text {*
2208
+ − 765
+ − 766
With the above definitions of @{text "REG"} and @{text "INJ"} we can show
+ − 767
how the proof is performed. The first step is always the application of
+ − 768
of the following lemma:
+ − 769
2231
+ − 770
@{term [display, indent=10] "[|A; A --> B; B = C; C = D|] ==> D"}
2208
+ − 771
+ − 772
With @{text A} instantiated to the original raw theorem,
+ − 773
@{text B} instantiated to @{text "REG(A)"},
+ − 774
@{text C} instantiated to @{text "INJ(REG(A))"},
+ − 775
and @{text D} instantiated to the statement of the lifted theorem.
+ − 776
The first assumption can be immediately discharged using the original
+ − 777
theorem and the three left subgoals are exactly the subgoals of regularization,
+ − 778
injection and cleaning. The three can be proved independently by the
+ − 779
framework and in case there are non-solved subgoals they can be left
+ − 780
to the user.
+ − 781
+ − 782
The injection and cleaning subgoals are always solved if the appropriate
+ − 783
respectfulness and preservation theorems are given. It is not the case
+ − 784
with regularization; sometimes a theorem given by the user does not
+ − 785
imply a regularized version and a stronger one needs to be proved. This
2242
+ − 786
is outside of the scope of the quotient package, so such obligations are
+ − 787
left to the user. Take a simple statement for integers @{term "0 \<noteq> 1"}.
+ − 788
It does not follow from the fact that @{term "\<not> (0, 0) = (1, 0)"} because
+ − 789
of regularization. The raw theorem only shows that particular items in the
+ − 790
equivalence classes are not equal. A more general statement saying that
+ − 791
the classes are not equal is necessary.
2208
+ − 792
*}
+ − 793
+ − 794
subsection {* Proving Regularization *}
+ − 795
+ − 796
text {*
1994
+ − 797
2209
+ − 798
Isabelle provides a set of \emph{mono} rules, that are used to split implications
2230
+ − 799
of similar statements into simpler implication subgoals. These are enhanced
2209
+ − 800
with special quotient theorem in the regularization goal. Below we only show
+ − 801
the versions for the universal quantifier. For the existential quantifier
2242
+ − 802
and abstraction they are analogous.
2199
+ − 803
2209
+ − 804
First, bounded universal quantifiers can be removed on the right:
2199
+ − 805
2231
+ − 806
@{thm [display, indent=10] ball_reg_right[no_vars]}
2206
+ − 807
2209
+ − 808
They can be removed anywhere if the relation is an equivalence relation:
+ − 809
2231
+ − 810
@{thm [display, indent=10] ball_reg_eqv[no_vars]}
2209
+ − 811
+ − 812
And finally it can be removed anywhere if @{term R2} is an equivalence relation, then:
2231
+ − 813
+ − 814
@{thm [display, indent=10] (concl) ball_reg_eqv_range[no_vars]}
2209
+ − 815
2242
+ − 816
The last theorem is new in comparison with Homeier's package. There the
2231
+ − 817
injection procedure would be used to prove goals with such shape, and there
2242
+ − 818
the equivalence assumption would be used. We use the above theorem directly
+ − 819
also for composed relations where the range type is a type for which we know an
2231
+ − 820
equivalence theorem. This allows separating regularization from injection.
2209
+ − 821
2206
+ − 822
*}
+ − 823
+ − 824
(*
2231
+ − 825
@{thm bex_reg_eqv_range[no_vars]}
2199
+ − 826
@{thm [display] bex_reg_left[no_vars]}
+ − 827
@{thm [display] bex1_bexeq_reg[no_vars]}
2206
+ − 828
@{thm [display] bex_reg_eqv[no_vars]}
2209
+ − 829
@{thm [display] babs_reg_eqv[no_vars]}
+ − 830
@{thm [display] babs_simp[no_vars]}
2206
+ − 831
*)
1994
+ − 832
+ − 833
subsection {* Injection *}
+ − 834
2199
+ − 835
text {*
2211
+ − 836
The injection proof starts with an equality between the regularized theorem
+ − 837
and the injected version. The proof again follows by the structure of the
2242
+ − 838
two terms, and is defined for a goal being a relation between these two terms.
2199
+ − 839
2211
+ − 840
\begin{itemize}
+ − 841
\item For two constants, an appropriate constant respectfullness assumption is used.
2242
+ − 842
\item For two variables, we use the assumptions proved in regularization.
2211
+ − 843
\item For two abstractions, they are eta-expanded and beta-reduced.
+ − 844
\end{itemize}
2199
+ − 845
2211
+ − 846
Otherwise the two terms are applications. There are two cases: If there is a REP/ABS
+ − 847
in the injected theorem we can use the theorem:
+ − 848
2243
+ − 849
@{thm [display, indent=10] rep_abs_rsp[no_vars]}
2199
+ − 850
2243
+ − 851
\noindent
2211
+ − 852
and continue the proof.
2199
+ − 853
2211
+ − 854
Otherwise we introduce an appropriate relation between the subterms and continue with
+ − 855
two subgoals using the lemma:
+ − 856
2243
+ − 857
@{thm [display, indent=10] apply_rsp[no_vars]}
2199
+ − 858
+ − 859
*}
+ − 860
1994
+ − 861
subsection {* Cleaning *}
+ − 862
2212
+ − 863
text {*
+ − 864
The @{text REG} and @{text INJ} functions have been defined in such a way
+ − 865
that establishing the goal theorem now consists only on rewriting the
+ − 866
injected theorem with the preservation theorems.
+ − 867
+ − 868
\begin{itemize}
+ − 869
\item First for lifted constants, their definitions are the preservation rules for
+ − 870
them.
+ − 871
\item For lambda abstractions lambda preservation establishes
+ − 872
the equality between the injected theorem and the goal. This allows both
+ − 873
abstraction and quantification over lifted types.
+ − 874
@{thm [display] lambda_prs[no_vars]}
+ − 875
\item Relations over lifted types are folded with:
+ − 876
@{thm [display] Quotient_rel_rep[no_vars]}
+ − 877
\item User given preservation theorems, that allow using higher level operations
+ − 878
and containers of types being lifted. An example may be
2243
+ − 879
@{thm [display] map_prs(1)[of R1 Abs1 Rep1 R2 Abs2 Rep2,no_vars]}
2212
+ − 880
\end{itemize}
+ − 881
+ − 882
Preservation of relations and user given constant preservation lemmas *}
1994
+ − 883
+ − 884
section {* Examples *}
+ − 885
2210
+ − 886
(* Mention why equivalence *)
2206
+ − 887
2210
+ − 888
text {*
+ − 889
2239
+ − 890
In this section we will show, a complete interaction with the quotient package
2240
+ − 891
for defining the type of integers by quotienting pairs of natural numbers and
+ − 892
lifting theorems to integers. Our quotient package is fully compatible with
+ − 893
Isabelle type classes, but for clarity we will not use them in this example.
+ − 894
In a larger formalization of integers using the type class mechanism would
+ − 895
provide many algebraic properties ``for free''.
2210
+ − 896
2240
+ − 897
A user of our quotient package first needs to define a relation on
+ − 898
the raw type, by which the quotienting will be performed. We give
+ − 899
the same integer relation as the one presented in the introduction:
+ − 900
+ − 901
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
+ − 902
\isacommand{fun}~~@{text "int_rel"}~~\isacommand{where}~~@{text "(m \<Colon> nat, n) int_rel (p, q) = (m + q = n + p)"}
2239
+ − 903
\end{isabelle}
2210
+ − 904
2239
+ − 905
\noindent
+ − 906
Next the quotient type is defined. This leaves a proof obligation that the
+ − 907
relation is an equivalence relation which is solved automatically using the
+ − 908
definitions:
2210
+ − 909
2240
+ − 910
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
+ − 911
\isacommand{quotient\_type}~~@{text "int"}~~\isacommand{=}~~@{text "(nat \<times> nat)"}~~\isacommand{/}~~@{text "int_rel"}
2239
+ − 912
\end{isabelle}
2210
+ − 913
2239
+ − 914
\noindent
2210
+ − 915
The user can then specify the constants on the quotient type:
+ − 916
2240
+ − 917
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
+ − 918
\begin{tabular}{@ {}l}
+ − 919
\isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
+ − 920
\isacommand{fun}~~@{text "plus_raw"}~~\isacommand{where}~~@{text "plus_raw (m :: nat, n) (p, q) = (m + p, n + q)"}\\
+ − 921
\isacommand{quotient\_definition}~~@{text "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)"}~~\isacommand{is}~~@{text "plus_raw"}\\
+ − 922
\end{tabular}
+ − 923
\end{isabelle}
2210
+ − 924
2240
+ − 925
\noindent
2210
+ − 926
Lets first take a simple theorem about addition on the raw level:
+ − 927
2240
+ − 928
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
+ − 929
\isacommand{lemma}~~@{text "plus_zero_raw: int_rel (plus_raw (0, 0) x) x"}
2240
+ − 930
\end{isabelle}
2210
+ − 931
2240
+ − 932
\noindent
2210
+ − 933
When the user tries to lift a theorem about integer addition, the respectfulness
+ − 934
proof obligation is left, so let us prove it first:
+ − 935
2240
+ − 936
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
+ − 937
\isacommand{lemma}~~@{text "[quot_respect]: (int_rel \<Longrightarrow> int_rel \<Longrightarrow> int_rel) plus_raw plus_raw"}
2240
+ − 938
\end{isabelle}
+ − 939
+ − 940
\noindent
2210
+ − 941
Can be proved automatically by the system just by unfolding the definition
2240
+ − 942
of @{text "op \<Longrightarrow>"}.
2230
+ − 943
Now the user can either prove a lifted lemma explicitly:
2210
+ − 944
2240
+ − 945
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
+ − 946
\isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting plus_zero_raw"}
+ − 947
\end{isabelle}
2210
+ − 948
2240
+ − 949
\noindent
2210
+ − 950
Or in this simple case use the automated translation mechanism:
+ − 951
2240
+ − 952
\begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
+ − 953
\isacommand{thm}~~@{text "plus_zero_raw[quot_lifted]"}
+ − 954
\end{isabelle}
2210
+ − 955
2240
+ − 956
\noindent
2210
+ − 957
obtaining the same result.
+ − 958
*}
2206
+ − 959
1978
+ − 960
section {* Related Work *}
+ − 961
+ − 962
text {*
+ − 963
\begin{itemize}
+ − 964
2152
+ − 965
\item Peter Homeier's package~\cite{Homeier05} (and related work from there)
2243
+ − 966
2152
+ − 967
\item John Harrison's one~\cite{harrison-thesis} is the first one to lift theorems
+ − 968
but only first order.
1978
+ − 969
2152
+ − 970
\item PVS~\cite{PVS:Interpretations}
+ − 971
\item MetaPRL~\cite{Nogin02}
2243
+ − 972
+ − 973
% \item Manually defined quotients in Isabelle/HOL Library (Markus's Quotient\_Type,
+ − 974
% Dixon's FSet, \ldots)
1978
+ − 975
+ − 976
\item Oscar Slotosch defines quotient-type automatically but no
2152
+ − 977
lifting~\cite{Slotosch97}.
1978
+ − 978
+ − 979
\item PER. And how to avoid it.
+ − 980
2152
+ − 981
\item Necessity of Hilbert Choice op and Larry's quotients~\cite{Paulson06}
1978
+ − 982
2152
+ − 983
\item Setoids in Coq and \cite{ChicliPS02}
1978
+ − 984
+ − 985
\end{itemize}
+ − 986
*}
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 987
2210
+ − 988
section {* Conclusion *}
+ − 989
2224
+ − 990
text {*
2237
+ − 991
+ − 992
+ − 993
The code of the quotient package described here is already included in the
+ − 994
standard distribution of Isabelle.\footnote{Avaiable from
+ − 995
\href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} It is
+ − 996
heavily used in Nominal Isabelle, which provides a convenient reasoning
+ − 997
infrastructure for programming language calculi involving binders. Earlier
+ − 998
versions of Nominal Isabelle have been used successfully in formalisations
+ − 999
of an equivalence checking algorithm for LF \cite{UrbanCheneyBerghofer08},
+ − 1000
Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for
+ − 1001
concurrency \cite{BengtsonParow09} and a strong normalisation result for
+ − 1002
cut-elimination in classical logic \cite{UrbanZhu08}.
+ − 1003
2224
+ − 1004
*}
+ − 1005
+ − 1006
+ − 1007
subsection {* Contributions *}
+ − 1008
+ − 1009
text {*
+ − 1010
We present the detailed lifting procedure, which was not shown before.
+ − 1011
+ − 1012
The quotient package presented in this paper has the following
+ − 1013
advantages over existing packages:
+ − 1014
\begin{itemize}
+ − 1015
+ − 1016
\item We define quotient composition, function map composition and
+ − 1017
relation map composition. This lets lifting polymorphic types with
+ − 1018
subtypes quotiented as well. We extend the notions of
+ − 1019
respectfulness and preservation to cope with quotient
+ − 1020
composition.
+ − 1021
+ − 1022
\item We allow lifting only some occurrences of quotiented
+ − 1023
types. Rsp/Prs extended. (used in nominal)
+ − 1024
+ − 1025
\item The quotient package is very modular. Definitions can be added
+ − 1026
separately, rsp and prs can be proved separately, Quotients and maps
+ − 1027
can be defined separately and theorems can
+ − 1028
be lifted on a need basis. (useful with type-classes).
+ − 1029
+ − 1030
\item Can be used both manually (attribute, separate tactics,
+ − 1031
rsp/prs databases) and programatically (automated definition of
+ − 1032
lifted constants, the rsp proof obligations and theorem statement
+ − 1033
translation according to given quotients).
+ − 1034
+ − 1035
\end{itemize}
+ − 1036
*}
+ − 1037
+ − 1038
2227
+ − 1039
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1040
(*<*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1041
end
1978
+ − 1042
(*>*)