1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1
(*<*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 2
theory Paper
2774
+ − 3
imports "Quotient"
+ − 4
"~~/src/HOL/Library/Quotient_Syntax"
2747
+ − 5
"~~/src/HOL/Library/LaTeXsugar"
+ − 6
"~~/src/HOL/Quotient_Examples/FSet"
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 7
begin
1994
+ − 8
2286
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 9
(****
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 10
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 11
** things to do for the next version
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 12
*
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 13
* - what are quot_thms?
2287
+ − 14
* - what do all preservation theorems look like,
+ − 15
in particular preservation for quotient
+ − 16
compositions
2333
+ − 17
- explain how Quotient R Abs Rep is proved (j-version)
+ − 18
- give an example where precise specification helps (core Haskell in nominal?)
2374
+ − 19
2414
+ − 20
- Mention Andreas Lochbiler in Acknowledgements and 'desceding'.
+ − 21
2286
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 22
*)
e7bc2ae30faf
added a few points that need to be looked at the next version of the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 23
1994
+ − 24
notation (latex output)
2258
+ − 25
rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
2269
+ − 26
pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
2457
+ − 27
implies (infix "\<longrightarrow>" 100) and
2258
+ − 28
"==>" (infix "\<Longrightarrow>" 100) and
2774
+ − 29
map_fun ("_ \<singlearr> _" 51) and
2444
+ − 30
fun_rel ("_ \<doublearr> _" 51) and
2217
+ − 31
list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
2549
+ − 32
empty_fset ("\<emptyset>") and
+ − 33
union_fset ("_ \<union> _") and
+ − 34
insert_fset ("{_} \<union> _") and
2221
+ − 35
Cons ("_::_") and
+ − 36
concat ("flat") and
2549
+ − 37
concat_fset ("\<Union>") and
2444
+ − 38
Quotient ("Quot _ _ _")
+ − 39
2217
+ − 40
1994
+ − 41
2182
+ − 42
ML {*
+ − 43
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
2274
+ − 44
2182
+ − 45
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
+ − 46
let
+ − 47
val concl =
+ − 48
Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
+ − 49
in
+ − 50
case concl of (_ $ l $ r) => proj (l, r)
+ − 51
| _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
+ − 52
end);
+ − 53
*}
2274
+ − 54
2182
+ − 55
setup {*
+ − 56
Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
+ − 57
Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
+ − 58
Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
+ − 59
*}
2274
+ − 60
2528
+ − 61
lemma insert_preserve2:
+ − 62
shows "((rep_fset ---> (map rep_fset \<circ> rep_fset) ---> (abs_fset \<circ> map abs_fset)) op #) =
+ − 63
(id ---> rep_fset ---> abs_fset) op #"
+ − 64
by (simp add: fun_eq_iff abs_o_rep[OF Quotient_fset] map_id Quotient_abs_rep[OF Quotient_fset])
+ − 65
2774
+ − 66
lemma list_all2_symp:
+ − 67
assumes a: "equivp R"
+ − 68
and b: "list_all2 R xs ys"
+ − 69
shows "list_all2 R ys xs"
+ − 70
using list_all2_lengthD[OF b] b
+ − 71
apply(induct xs ys rule: list_induct2)
+ − 72
apply(auto intro: equivp_symp[OF a])
+ − 73
done
+ − 74
2529
+ − 75
lemma concat_rsp_unfolded:
+ − 76
"\<lbrakk>list_all2 list_eq a ba; list_eq ba bb; list_all2 list_eq bb b\<rbrakk> \<Longrightarrow> list_eq (concat a) (concat b)"
+ − 77
proof -
+ − 78
fix a b ba bb
+ − 79
assume a: "list_all2 list_eq a ba"
+ − 80
assume b: "list_eq ba bb"
+ − 81
assume c: "list_all2 list_eq bb b"
+ − 82
have "\<forall>x. (\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
+ − 83
fix x
+ − 84
show "(\<exists>xa\<in>set a. x \<in> set xa) = (\<exists>xa\<in>set b. x \<in> set xa)" proof
+ − 85
assume d: "\<exists>xa\<in>set a. x \<in> set xa"
+ − 86
show "\<exists>xa\<in>set b. x \<in> set xa" by (rule concat_rsp_pre[OF a b c d])
+ − 87
next
+ − 88
assume e: "\<exists>xa\<in>set b. x \<in> set xa"
+ − 89
have a': "list_all2 list_eq ba a" by (rule list_all2_symp[OF list_eq_equivp, OF a])
+ − 90
have b': "list_eq bb ba" by (rule equivp_symp[OF list_eq_equivp, OF b])
+ − 91
have c': "list_all2 list_eq b bb" by (rule list_all2_symp[OF list_eq_equivp, OF c])
+ − 92
show "\<exists>xa\<in>set a. x \<in> set xa" by (rule concat_rsp_pre[OF c' b' a' e])
+ − 93
qed
+ − 94
qed
+ − 95
then show "list_eq (concat a) (concat b)" by auto
+ − 96
qed
+ − 97
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 98
(*>*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 99
2227
+ − 100
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 101
section {* Introduction *}
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 102
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 103
text {*
2443
+ − 104
\noindent
2445
+ − 105
One might think quotients have been studied to death, but in the context of
+ − 106
theorem provers many questions concerning them are far from settled. In
+ − 107
this paper we address the question of how to establish a convenient reasoning
+ − 108
infrastructure
2811
+ − 109
for quotient constructions in the Isabelle/HOL
2445
+ − 110
theorem prover. Higher-Order Logic (HOL) consists
2443
+ − 111
of a small number of axioms and inference rules over a simply-typed
+ − 112
term-language. Safe reasoning in HOL is ensured by two very restricted
+ − 113
mechanisms for extending the logic: one is the definition of new constants
+ − 114
in terms of existing ones; the other is the introduction of new types by
2445
+ − 115
identifying non-empty subsets in existing types. Previous work has shown how
2443
+ − 116
to use both mechanisms for dealing with quotient constructions in HOL (see
+ − 117
\cite{Homeier05,Paulson06}). For example the integers in Isabelle/HOL are
+ − 118
constructed by a quotient construction over the type @{typ "nat \<times> nat"} and
+ − 119
the equivalence relation
2103
+ − 120
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 121
2443
+ − 122
\begin{isabelle}\ \ \ \ \ %%%
2237
+ − 123
@{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}
+ − 124
\end{isabelle}
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 125
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 126
\noindent
2554
+ − 127
This constructions yields the new type @{typ int}, and definitions for @{text
2220
+ − 128
"0"} and @{text "1"} of type @{typ int} can be given in terms of pairs of
+ − 129
natural numbers (namely @{text "(0, 0)"} and @{text "(1, 0)"}). Operations
+ − 130
such as @{text "add"} with type @{typ "int \<Rightarrow> int \<Rightarrow> int"} can be defined in
+ − 131
terms of operations on pairs of natural numbers (namely @{text
2224
+ − 132
"add_pair (n\<^isub>1, m\<^isub>1) (n\<^isub>2,
2222
+ − 133
m\<^isub>2) \<equiv> (n\<^isub>1 + n\<^isub>2, m\<^isub>1 + m\<^isub>2)"}).
2554
+ − 134
Similarly one can construct %%the type of
+ − 135
finite sets, written @{term "\<alpha> fset"},
2223
+ − 136
by quotienting the type @{text "\<alpha> list"} according to the equivalence relation
2103
+ − 137
2443
+ − 138
\begin{isabelle}\ \ \ \ \ %%%
2237
+ − 139
@{text "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}\hfill\numbered{listequiv}
+ − 140
\end{isabelle}
2103
+ − 141
+ − 142
\noindent
2224
+ − 143
which states that two lists are equivalent if every element in one list is
+ − 144
also member in the other. The empty finite set, written @{term "{||}"}, can
+ − 145
then be defined as the empty list and the union of two finite sets, written
+ − 146
@{text "\<union>"}, as list append.
2220
+ − 147
2269
+ − 148
Quotients are important in a variety of areas, but they are really ubiquitous in
2258
+ − 149
the area of reasoning about programming language calculi. A simple example
+ − 150
is the lambda-calculus, whose raw terms are defined as
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 151
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 152
\begin{isabelle}\ \ \ \ \ %%%
2554
+ − 153
@{text "t ::= x | t t | \<lambda>x.t"}%\hfill\numbered{lambda}
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 154
\end{isabelle}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 155
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 156
\noindent
2256
+ − 157
The problem with this definition arises, for instance, when one attempts to
+ − 158
prove formally the substitution lemma \cite{Barendregt81} by induction
2222
+ − 159
over the structure of terms. This can be fiendishly complicated (see
2220
+ − 160
\cite[Pages 94--104]{CurryFeys58} for some ``rough'' sketches of a proof
2226
+ − 161
about raw lambda-terms). In contrast, if we reason about
2220
+ − 162
$\alpha$-equated lambda-terms, that means terms quotient according to
2223
+ − 163
$\alpha$-equivalence, then the reasoning infrastructure provided,
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 164
for example, by Nominal Isabelle %%\cite{UrbanKaliszyk11}
2455
+ − 165
makes the formal
2222
+ − 166
proof of the substitution lemma almost trivial.
2220
+ − 167
+ − 168
The difficulty is that in order to be able to reason about integers, finite
2221
+ − 169
sets or $\alpha$-equated lambda-terms one needs to establish a reasoning
2220
+ − 170
infrastructure by transferring, or \emph{lifting}, definitions and theorems
2226
+ − 171
from the raw type @{typ "nat \<times> nat"} to the quotient type @{typ int}
2220
+ − 172
(similarly for finite sets and $\alpha$-equated lambda-terms). This lifting
2222
+ − 173
usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.
2443
+ − 174
In principle it is feasible to do this work manually, if one has only a few quotient
2247
+ − 175
constructions at hand. But if they have to be done over and over again, as in
2222
+ − 176
Nominal Isabelle, then manual reasoning is not an option.
2221
+ − 177
2223
+ − 178
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
+ − 179
and automate the reasoning as much as possible. In the
2223
+ − 180
context of HOL, there have been a few quotient packages already
2234
+ − 181
\cite{harrison-thesis,Slotosch97}. The most notable one is by Homeier
2223
+ − 182
\cite{Homeier05} implemented in HOL4. The fundamental construction these
+ − 183
quotient packages perform can be illustrated by the following picture:
2217
+ − 184
2417
+ − 185
%%% FIXME: Referee 1 says:
+ − 186
%%% Diagram is unclear. Firstly, isn't an existing type a "set (not sets) of raw elements"?
+ − 187
%%% Secondly, isn't the _set of_ equivalence classes mapped to and from the new type?
+ − 188
%%% Thirdly, what do the words "non-empty subset" refer to ?
+ − 189
2423
+ − 190
%%% Cezary: I like the diagram, maybe 'new type' could be outside, but otherwise
+ − 191
%%% I wouldn't change it.
+ − 192
2217
+ − 193
\begin{center}
2554
+ − 194
\mbox{}\hspace{20mm}\begin{tikzpicture}[scale=0.9]
2220
+ − 195
%%\draw[step=2mm] (-4,-1) grid (4,1);
+ − 196
+ − 197
\draw[very thick] (0.7,0.3) circle (4.85mm);
+ − 198
\draw[rounded corners=1mm, very thick] ( 0.0,-0.9) rectangle ( 1.8, 0.9);
+ − 199
\draw[rounded corners=1mm, very thick] (-1.95,0.8) rectangle (-2.9,-0.195);
+ − 200
+ − 201
\draw (-2.0, 0.8) -- (0.7,0.8);
+ − 202
\draw (-2.0,-0.195) -- (0.7,-0.195);
2103
+ − 203
2220
+ − 204
\draw ( 0.7, 0.23) node {\begin{tabular}{@ {}c@ {}}equiv-\\[-1mm]clas.\end{tabular}};
+ − 205
\draw (-2.45, 0.35) node {\begin{tabular}{@ {}c@ {}}new\\[-1mm]type\end{tabular}};
+ − 206
\draw (1.8, 0.35) node[right=-0.1mm]
2224
+ − 207
{\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw elements)\end{tabular}};
2220
+ − 208
\draw (0.9, -0.55) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
+ − 209
+ − 210
\draw[->, very thick] (-1.8, 0.36) -- (-0.1,0.36);
+ − 211
\draw[<-, very thick] (-1.8, 0.16) -- (-0.1,0.16);
2221
+ − 212
\draw (-0.95, 0.26) node[above=0.4mm] {@{text Rep}};
+ − 213
\draw (-0.95, 0.26) node[below=0.4mm] {@{text Abs}};
2220
+ − 214
\end{tikzpicture}
+ − 215
\end{center}
2217
+ − 216
2221
+ − 217
\noindent
2247
+ − 218
The starting point is an existing type, to which we refer as the
2554
+ − 219
\emph{raw type} and over which an equivalence relation is given by the user.
+ − 220
With this input the package introduces a new type, to which we
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 221
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
+ − 222
\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
+ − 223
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
+ − 224
the functions @{text Abs} and @{text Rep} need to be derived from them. We
2269
+ − 225
will show the details later. } They relate elements in the
2554
+ − 226
existing type to elements in the new type, % and vice versa,
+ − 227
and can be uniquely
2256
+ − 228
identified by their quotient type. For example for the integer quotient construction
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 229
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
+ − 230
2224
+ − 231
2443
+ − 232
\begin{isabelle}\ \ \ \ \ %%%
2234
+ − 233
@{text "Abs :: nat \<times> nat \<Rightarrow> int"}\hspace{10mm}@{text "Rep :: int \<Rightarrow> nat \<times> nat"}
2224
+ − 234
\end{isabelle}
+ − 235
+ − 236
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 237
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
+ − 238
typing information is important.
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 239
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 240
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
+ − 241
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
+ − 242
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
+ − 243
"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
+ − 244
2221
+ − 245
2443
+ − 246
\begin{isabelle}\ \ \ \ \ %%%
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 247
@{text "0 \<equiv> Abs_int (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs_int (1, 0)"}
2221
+ − 248
\end{isabelle}
+ − 249
+ − 250
\noindent
2224
+ − 251
Slightly more complicated is the definition of @{text "add"} having type
2222
+ − 252
@{typ "int \<Rightarrow> int \<Rightarrow> int"}. Its definition is as follows
2221
+ − 253
2443
+ − 254
\begin{isabelle}\ \ \ \ \ %%%
2258
+ − 255
@{text "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
+ − 256
\hfill\numbered{adddef}
+ − 257
\end{isabelle}
+ − 258
2221
+ − 259
\noindent
2224
+ − 260
where we take the representation of the arguments @{text n} and @{text m},
2226
+ − 261
add them according to the function @{text "add_pair"} and then take the
2221
+ − 262
abstraction of the result. This is all straightforward and the existing
2412
+ − 263
quotient packages can deal with such definitions. But what is surprising is
2223
+ − 264
that none of them can deal with slightly more complicated definitions involving
2221
+ − 265
\emph{compositions} of quotients. Such compositions are needed for example
2247
+ − 266
in case of quotienting lists to yield finite sets and the operator that
2226
+ − 267
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
+ − 268
2443
+ − 269
\begin{isabelle}\ \ \ \ \ %%%
+ − 270
@{thm concat.simps(1)[THEN eq_reflection]}\hspace{10mm}
+ − 271
@{thm concat.simps(2)[THEN eq_reflection, no_vars]}
+ − 272
\end{isabelle}
2183
+ − 273
2103
+ − 274
\noindent
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 275
where @{text "@"} is the usual
2554
+ − 276
list append. We expect that the corresponding
2527
40187684fc16
fixed the typo in the abstract and the problem with append (the type of map_k
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 277
operator on finite sets, written @{term "fconcat"},
2248
+ − 278
builds finite unions of finite sets:
2188
+ − 279
2443
+ − 280
\begin{isabelle}\ \ \ \ \ %%%
2549
+ − 281
@{thm concat_empty_fset[THEN eq_reflection, no_vars]}\hspace{10mm}
+ − 282
@{thm concat_insert_fset[THEN eq_reflection, no_vars]}
2443
+ − 283
\end{isabelle}
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 284
2103
+ − 285
\noindent
2256
+ − 286
The quotient package should automatically 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
+ − 287
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
+ − 288
that the method used in the existing quotient
2247
+ − 289
packages of just taking the representation of the arguments and then taking
2258
+ − 290
the abstraction of the result is \emph{not} enough. The reason is that in case
2224
+ − 291
of @{text "\<Union>"} we obtain the incorrect definition
+ − 292
2443
+ − 293
\begin{isabelle}\ \ \ \ \ %%%
+ − 294
@{text "\<Union> S \<equiv> Abs_fset (flat (Rep_fset S))"}
+ − 295
\end{isabelle}
2221
+ − 296
2224
+ − 297
\noindent
+ − 298
where the right-hand side is not even typable! This problem can be remedied in the
+ − 299
existing quotient packages by introducing an intermediate step and reasoning
2226
+ − 300
about flattening of lists of finite sets. However, this remedy is rather
2224
+ − 301
cumbersome and inelegant in light of our work, which can deal with such
+ − 302
definitions directly. The solution is that we need to build aggregate
+ − 303
representation and abstraction functions, which in case of @{text "\<Union>"}
2554
+ − 304
generate the %%%following
+ − 305
definition
2224
+ − 306
2443
+ − 307
\begin{isabelle}\ \ \ \ \ %%%
+ − 308
@{text "\<Union> S \<equiv> Abs_fset (flat ((map_list Rep_fset \<circ> Rep_fset) S))"}
+ − 309
\end{isabelle}
2221
+ − 310
+ − 311
\noindent
2367
+ − 312
where @{term map_list} is the usual mapping function for lists. In this paper we
2224
+ − 313
will present a formal definition of our aggregate abstraction and
2223
+ − 314
representation functions (this definition was omitted in \cite{Homeier05}).
2224
+ − 315
They generate definitions, like the one above for @{text "\<Union>"},
2226
+ − 316
according to the type of the raw constant and the type
2224
+ − 317
of the quotient constant. This means we also have to extend the notions
2237
+ − 318
of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
2231
+ − 319
from Homeier \cite{Homeier05}.
2223
+ − 320
2445
+ − 321
In addition we are able to clearly specify what is involved
2256
+ − 322
in the lifting process (this was only hinted at in \cite{Homeier05} and
2445
+ − 323
implemented as a ``rough recipe'' in ML-code). A pleasing side-result
+ − 324
is that our procedure for lifting theorems is completely deterministic
+ − 325
following the structure of the theorem being lifted and the theorem
+ − 326
on the quotient level. Space constraints, unfortunately, allow us to only
+ − 327
sketch this part of our work in Section 5 and we defer the reader to a longer
+ − 328
version for the details. However, we will give in Section 3 and 4 all
+ − 329
definitions that specify the input and output data of our three-step
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 330
lifting procedure. Appendix A gives an example how our quotient
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 331
package works in practise.
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 332
*}
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 333
2443
+ − 334
section {* Preliminaries and General\\ Quotients\label{sec:prelims} *}
1978
+ − 335
+ − 336
text {*
2443
+ − 337
\noindent
2445
+ − 338
We will give in this section a crude overview of HOL and describe the main
2274
+ − 339
definitions given by Homeier for quotients \cite{Homeier05}.
2258
+ − 340
2274
+ − 341
At its core, HOL is based on a simply-typed term language, where types are
2269
+ − 342
recorded in Church-style fashion (that means, we can always infer the type of
2258
+ − 343
a term and its subterms without any additional information). The grammars
2443
+ − 344
for types and terms are
2256
+ − 345
2443
+ − 346
\begin{isabelle}\ \ \ \ \ %%%
+ − 347
\begin{tabular}{@ {}c@ {\hspace{10mm}}c@ {}}
+ − 348
@{text "\<sigma>, \<tau> ::= \<alpha> | (\<sigma>,\<dots>, \<sigma>) \<kappa>"} &
+ − 349
@{text "t, s ::= x\<^isup>\<sigma> | c\<^isup>\<sigma> | t t | \<lambda>x\<^isup>\<sigma>. t"}\\
2256
+ − 350
\end{tabular}
+ − 351
\end{isabelle}
+ − 352
+ − 353
\noindent
2443
+ − 354
with types being either type variables or type constructors and terms
+ − 355
being variables, constants, applications or abstractions.
2256
+ − 356
We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
2258
+ − 357
@{text "\<sigma>s"} to stand for collections of type variables and types,
+ − 358
respectively. The type of a term is often made explicit by writing @{text
2274
+ − 359
"t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
+ − 360
type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 361
constants; for example, a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
2287
+ − 362
bool"}, and the identity function with type @{text "id :: \<sigma> \<Rightarrow> \<sigma>"} is
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 363
defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}.
2256
+ − 364
2258
+ − 365
An important point to note is that theorems in HOL can be seen as a subset
2287
+ − 366
of terms that are constructed specially (namely through axioms and proof
2269
+ − 367
rules). As a result we are able to define automatic proof
2258
+ − 368
procedures showing that one theorem implies another by decomposing the term
+ − 369
underlying the first theorem.
+ − 370
2333
+ − 371
Like Homeier's, our work relies on map-functions defined for every type
2367
+ − 372
constructor taking some arguments, for example @{text map_list} for lists. Homeier
2274
+ − 373
describes in \cite{Homeier05} map-functions for products, sums, options and
+ − 374
also the following map for function types
2258
+ − 375
2443
+ − 376
\begin{isabelle}\ \ \ \ \ %%%
2774
+ − 377
@{thm map_fun_def[no_vars, THEN eq_reflection]}
2443
+ − 378
\end{isabelle}
2258
+ − 379
+ − 380
\noindent
+ − 381
Using this map-function, we can give the following, equivalent, but more
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 382
uniform definition for @{text add} shown in \eqref{adddef}:
2256
+ − 383
2443
+ − 384
\begin{isabelle}\ \ \ \ \ %%%
+ − 385
@{text "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
+ − 386
\end{isabelle}
2182
+ − 387
2258
+ − 388
\noindent
2274
+ − 389
Using extensionality and unfolding the definition of @{text "\<singlearr>"},
+ − 390
we can get back to \eqref{adddef}.
+ − 391
In what follows we shall use the convention to write @{text "map_\<kappa>"} for a map-function
2549
+ − 392
of the type-constructor @{text \<kappa>}.
+ − 393
%% a general type for map all types is difficult to give (algebraic types are
+ − 394
%% easy, but for example the function type is not algebraic
+ − 395
%For a type @{text \<kappa>} with arguments @{text "\<alpha>\<^isub>1\<^isub>\<dots>\<^isub>n"} the
+ − 396
%type of the function @{text "map_\<kappa>"} has to be @{text "\<alpha>\<^isub>1\<Rightarrow>\<dots>\<Rightarrow>\<alpha>\<^isub>n\<Rightarrow>\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>n \<kappa>"}.
+ − 397
%For example @{text "map_list"}
+ − 398
%has to have the type @{text "\<alpha>\<Rightarrow>\<alpha> list"}.
2413
+ − 399
In our implementation we maintain
2274
+ − 400
a database of these map-functions that can be dynamically extended.
2258
+ − 401
+ − 402
It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
+ − 403
which define equivalence relations in terms of constituent equivalence
+ − 404
relations. For example given two equivalence relations @{text "R\<^isub>1"}
+ − 405
and @{text "R\<^isub>2"}, we can define an equivalence relations over
2554
+ − 406
products as %% follows
2443
+ − 407
+ − 408
\begin{isabelle}\ \ \ \ \ %%%
+ − 409
@{text "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
+ − 410
\end{isabelle}
1978
+ − 411
2258
+ − 412
\noindent
2269
+ − 413
Homeier gives also the following operator for defining equivalence
+ − 414
relations over function types
2258
+ − 415
%
2443
+ − 416
\begin{isabelle}\ \ \ \ \ %%%
2274
+ − 417
@{thm fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
+ − 418
\hfill\numbered{relfun}
+ − 419
\end{isabelle}
+ − 420
+ − 421
\noindent
2319
+ − 422
In the context of quotients, the following two notions from \cite{Homeier05}
+ − 423
are needed later on.
2274
+ − 424
+ − 425
\begin{definition}[Respects]\label{def:respects}
+ − 426
An element @{text "x"} respects a relation @{text "R"} provided @{text "R x x"}.
+ − 427
\end{definition}
+ − 428
2444
+ − 429
\begin{definition}[Bounded $\forall$ and $\lambda$]\label{def:babs}
2274
+ − 430
@{text "\<forall>x \<in> S. P x"} holds if for all @{text x}, @{text "x \<in> S"} implies @{text "P x"};
+ − 431
and @{text "(\<lambda>x \<in> S. f x) = f x"} provided @{text "x \<in> S"}.
+ − 432
\end{definition}
2258
+ − 433
+ − 434
The central definition in Homeier's work \cite{Homeier05} relates equivalence
+ − 435
relations, abstraction and representation functions:
+ − 436
+ − 437
\begin{definition}[Quotient Types]
+ − 438
Given a relation $R$, an abstraction function $Abs$
+ − 439
and a representation function $Rep$, the predicate @{term "Quotient R Abs Rep"}
2413
+ − 440
holds if and only if
2444
+ − 441
\begin{isabelle}\ \ \ \ \ %%%%
+ − 442
\begin{tabular}{rl}
+ − 443
(i) & \begin{isa}@{thm (rhs1) Quotient_def[of "R", no_vars]}\end{isa}\\
+ − 444
(ii) & \begin{isa}@{thm (rhs2) Quotient_def[of "R", no_vars]}\end{isa}\\
+ − 445
(iii) & \begin{isa}@{thm (rhs3) Quotient_def[of "R", no_vars]}\end{isa}\\
+ − 446
\end{tabular}
+ − 447
\end{isabelle}
2182
+ − 448
\end{definition}
+ − 449
2258
+ − 450
\noindent
2444
+ − 451
The value of this definition lies in the fact that validity of @{term "Quotient R Abs Rep"} can
+ − 452
often be proved in terms of the validity of @{term "Quot"} over the constituent
2269
+ − 453
types of @{text "R"}, @{text Abs} and @{text Rep}.
2258
+ − 454
For example Homeier proves the following property for higher-order quotient
+ − 455
types:
+ − 456
2269
+ − 457
\begin{proposition}\label{funquot}
2444
+ − 458
\begin{isa}
2258
+ − 459
@{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2"
+ − 460
and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
2444
+ − 461
\end{isa}
2258
+ − 462
\end{proposition}
2195
+ − 463
2258
+ − 464
\noindent
2274
+ − 465
As a result, Homeier is able to build an automatic prover that can nearly
2444
+ − 466
always discharge a proof obligation involving @{text "Quot"}. Our quotient
2269
+ − 467
package makes heavy
+ − 468
use of this part of Homeier's work including an extension
2445
+ − 469
for dealing with \emph{conjugations} of equivalence relations\footnote{That are
+ − 470
symmetric by definition.} defined as follows:
2234
+ − 471
2417
+ − 472
%%% FIXME Referee 2 claims that composition-of-relations means OO, and this is also
+ − 473
%%% what wikipedia says. Any idea for a different name? Conjugation of Relations?
+ − 474
2444
+ − 475
\begin{definition}%%[Composition of Relations]
2258
+ − 476
@{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
2319
+ − 477
composition defined by
+ − 478
@{thm (concl) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}
2366
+ − 479
holds if and only if there exists a @{text y} such that @{thm (prem 1) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]} and
2274
+ − 480
@{thm (prem 2) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}.
2258
+ − 481
\end{definition}
2237
+ − 482
2258
+ − 483
\noindent
2413
+ − 484
Unfortunately a general quotient theorem for @{text "\<circ>\<circ>\<circ>"}, analogous to the one
+ − 485
for @{text "\<singlearr>"} given in Proposition \ref{funquot}, would not be true
+ − 486
in general. It cannot even be stated inside HOL, because of restrictions on types.
+ − 487
However, we can prove specific instances of a
2282
+ − 488
quotient theorem for composing particular quotient relations.
2287
+ − 489
For example, to lift theorems involving @{term flat} the quotient theorem for
+ − 490
composing @{text "\<approx>\<^bsub>list\<^esub>"} will be necessary: given @{term "Quotient R Abs Rep"}
+ − 491
with @{text R} being an equivalence relation, then
2282
+ − 492
2443
+ − 493
\begin{isabelle}\ \ \ \ \ %%%
2444
+ − 494
\begin{tabular}{r@ {\hspace{1mm}}l}
+ − 495
@{text "Quot"} & @{text "(rel_list R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>)"}\\
2554
+ − 496
& @{text "(Abs_fset \<circ> map_list Abs)"} @{text "(map_list Rep \<circ> Rep_fset)"}\\
2444
+ − 497
\end{tabular}
2443
+ − 498
\end{isabelle}
2195
+ − 499
*}
+ − 500
2444
+ − 501
section {* Quotient Types and Quotient\\ Definitions\label{sec:type} *}
1978
+ − 502
2234
+ − 503
text {*
2445
+ − 504
\noindent
2247
+ − 505
The first step in a quotient construction 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
+ − 506
type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
2247
+ − 507
defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
2269
+ − 508
relation must be @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
+ − 509
the quotient type declaration is therefore
2234
+ − 510
2443
+ − 511
\begin{isabelle}\ \ \ \ \ %%%
2269
+ − 512
\isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}\hfill\numbered{typedecl}
2235
+ − 513
\end{isabelle}
+ − 514
+ − 515
\noindent
2445
+ − 516
and a proof that @{text "R"} is indeed an equivalence relation. The @{text "\<alpha>s"}
+ − 517
indicate the arity of the new type and the type-variables of @{text "\<sigma>"} can only
+ − 518
be contained in @{text "\<alpha>s"}. Two concrete
2237
+ − 519
examples are
+ − 520
+ − 521
2443
+ − 522
\begin{isabelle}\ \ \ \ \ %%%
2237
+ − 523
\begin{tabular}{@ {}l}
+ − 524
\isacommand{quotient\_type}~~@{text "int = nat \<times> nat / \<approx>\<^bsub>nat \<times> nat\<^esub>"}\\
+ − 525
\isacommand{quotient\_type}~~@{text "\<alpha> fset = \<alpha> list / \<approx>\<^bsub>list\<^esub>"}
+ − 526
\end{tabular}
+ − 527
\end{isabelle}
+ − 528
+ − 529
\noindent
+ − 530
which introduce the type of integers and of finite sets using the
+ − 531
equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
2269
+ − 532
"\<approx>\<^bsub>list\<^esub>"} defined in \eqref{natpairequiv} and
2247
+ − 533
\eqref{listequiv}, respectively (the proofs about being equivalence
2811
+ − 534
relations are omitted). Given this data, we define for declarations shown in
2274
+ − 535
\eqref{typedecl} the quotient types internally as
2234
+ − 536
2443
+ − 537
\begin{isabelle}\ \ \ \ \ %%%
2234
+ − 538
\isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
+ − 539
\end{isabelle}
+ − 540
+ − 541
\noindent
2247
+ − 542
where the right-hand side is the (non-empty) set of equivalence classes of
2274
+ − 543
@{text "R"}. The constraint in this declaration is that the type variables
2237
+ − 544
in the raw type @{text "\<sigma>"} must be included in the type variables @{text
2274
+ − 545
"\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will then provide us with the following
2269
+ − 546
abstraction and representation functions
2182
+ − 547
2443
+ − 548
\begin{isabelle}\ \ \ \ \ %%%
2234
+ − 549
@{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"}
+ − 550
\end{isabelle}
+ − 551
2235
+ − 552
\noindent
2269
+ − 553
As can be seen from the type, they relate the new quotient type and equivalence classes of the raw
2235
+ − 554
type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
+ − 555
to work with the following derived abstraction and representation functions
+ − 556
2443
+ − 557
\begin{isabelle}\ \ \ \ \ %%%
2234
+ − 558
@{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)"}
+ − 559
\end{isabelle}
+ − 560
+ − 561
\noindent
2235
+ − 562
on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"} in the
2237
+ − 563
definition of @{text "Rep_\<kappa>\<^isub>q"}. These derived notions relate the
+ − 564
quotient type and the raw type directly, as can be seen from their type,
+ − 565
namely @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"},
+ − 566
respectively. Given that @{text "R"} is an equivalence relation, the
2274
+ − 567
following property holds for every quotient type
+ − 568
(for the proof see \cite{Homeier05}).
2237
+ − 569
2258
+ − 570
\begin{proposition}
2445
+ − 571
\begin{isa}@{term "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}.\end{isa}
2258
+ − 572
\end{proposition}
2234
+ − 573
2247
+ − 574
The next step in a quotient construction is to introduce definitions of new constants
+ − 575
involving the quotient type. These definitions need to be given in terms of concepts
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 576
of the raw type (remember this is the only way how to extend HOL
2269
+ − 577
with new definitions). For the user the visible part of such definitions is the declaration
2235
+ − 578
2443
+ − 579
\begin{isabelle}\ \ \ \ \ %%%
2237
+ − 580
\isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
2235
+ − 581
\end{isabelle}
+ − 582
2237
+ − 583
\noindent
+ − 584
where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
+ − 585
and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
+ − 586
given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ
2269
+ − 587
in places where a quotient and raw type is involved). Two concrete examples are
2188
+ − 588
2443
+ − 589
\begin{isabelle}\ \ \ \ \ %%%
2237
+ − 590
\begin{tabular}{@ {}l}
+ − 591
\isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
+ − 592
\isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
+ − 593
\isacommand{is}~~@{text "flat"}
+ − 594
\end{tabular}
+ − 595
\end{isabelle}
+ − 596
+ − 597
\noindent
+ − 598
The first one declares zero for integers and the second the operator for
2274
+ − 599
building unions of finite sets (@{text "flat"} having the type
+ − 600
@{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"}).
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 601
2413
+ − 602
From such declarations given by the user, the quotient package needs to derive proper
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 603
definitions using @{text "Abs"} and @{text "Rep"}. The data we rely on is the given quotient type
2247
+ − 604
@{text "\<tau>"} and the raw type @{text "\<sigma>"}. They allow us to define \emph{aggregate
+ − 605
abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 606
\<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we shall give below. The idea behind
2274
+ − 607
these two functions is to simultaneously descend into the raw types @{text \<sigma>} and
2247
+ − 608
quotient types @{text \<tau>}, and generate the appropriate
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 609
@{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
2269
+ − 610
we generate just the identity whenever the types are equal. On the ``way'' down,
+ − 611
however we might have to use map-functions to let @{text Abs} and @{text Rep} act
2274
+ − 612
over the appropriate types. In what follows we use the short-hand notation
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 613
@{text "ABS (\<sigma>s, \<tau>s)"} to mean @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1)\<dots>ABS (\<sigma>\<^isub>n, \<tau>\<^isub>n)"}; similarly
2274
+ − 614
for @{text REP}.
+ − 615
%
2227
+ − 616
\begin{center}
2252
+ − 617
\hfill
2444
+ − 618
\begin{tabular}{@ {\hspace{2mm}}l@ {}}
+ − 619
\multicolumn{1}{@ {}l}{equal types:}\\
2554
+ − 620
@{text "ABS (\<sigma>, \<sigma>)"} $\dn$ @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\hspace{5mm}%\\
2444
+ − 621
@{text "REP (\<sigma>, \<sigma>)"} $\dn$ @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\smallskip\\
+ − 622
\multicolumn{1}{@ {}l}{function types:}\\
+ − 623
@{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)"}\\
+ − 624
@{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\\
+ − 625
\multicolumn{1}{@ {}l}{equal type constructors:}\\
+ − 626
@{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "map_\<kappa> (ABS (\<sigma>s, \<tau>s))"}\\
+ − 627
@{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "map_\<kappa> (REP (\<sigma>s, \<tau>s))"}\smallskip\\
+ − 628
\multicolumn{1}{@ {}l}{unequal type constructors:}\\
+ − 629
@{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)))"}\\
+ − 630
@{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"}
2247
+ − 631
\end{tabular}\hfill\numbered{ABSREP}
2227
+ − 632
\end{center}
2234
+ − 633
%
2232
+ − 634
\noindent
2445
+ − 635
In the last two clauses are subtle. We rely in them on the fact that the type @{text "\<alpha>s
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 636
\<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
2237
+ − 637
@{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
2445
+ − 638
list"}). This data is given by declarations shown in \eqref{typedecl}.
+ − 639
The quotient construction ensures that the type variables in @{text
2367
+ − 640
"\<rho>s \<kappa>"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
+ − 641
substitutions for the @{text "\<alpha>s"} when matching @{text "\<sigma>s \<kappa>"} against
2445
+ − 642
@{text "\<rho>s \<kappa>"}. This calculation determines what are the types in place
+ − 643
of the type variables @{text "\<alpha>s"} in the instance of
+ − 644
quotient type @{text "\<alpha>s \<kappa>\<^isub>q"}---namely @{text "\<tau>s"}, and the corresponding
+ − 645
types in place of the @{text "\<alpha>s"} in the raw type @{text "\<rho>s \<kappa>"}---namely @{text "\<sigma>s'"}. The
2237
+ − 646
function @{text "MAP"} calculates an \emph{aggregate map-function} for a raw
+ − 647
type as follows:
+ − 648
%
2227
+ − 649
\begin{center}
2444
+ − 650
\begin{tabular}{r@ {\hspace{1mm}}c@ {\hspace{1mm}}l}
2237
+ − 651
@{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
+ − 652
@{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id :: \<kappa> \<Rightarrow> \<kappa>"}\\
2232
+ − 653
@{text "MAP' (\<sigma>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (MAP'(\<sigma>s))"}\smallskip\\
2233
+ − 654
@{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
+ − 655
\end{tabular}
2227
+ − 656
\end{center}
2237
+ − 657
%
2232
+ − 658
\noindent
2366
+ − 659
In this definition we rely on the fact that in the first clause 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
+ − 660
term variables @{text a}. In the last clause we build an abstraction over all
2274
+ − 661
term-variables of the map-function generated by the auxiliary function
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 662
@{text "MAP'"}.
2274
+ − 663
The need for aggregate map-functions can be seen in cases where we build quotients,
2247
+ − 664
say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types, say @{text "(\<alpha> list) \<times> \<beta>"}.
+ − 665
In this case @{text MAP} generates the
+ − 666
aggregate map-function:
2232
+ − 667
2421
+ − 668
%%% FIXME: Reviewer 2 asks: last two lines defining ABS and REP for
+ − 669
%%% unequal type constructors: How are the $\varrho$s defined? The
+ − 670
%%% following paragraph mentions them, but this paragraph is unclear,
+ − 671
%%% since it then mentions $\alpha$s, which do not seem to be defined
+ − 672
%%% either. As a result, I do not understand the first two sentences
+ − 673
%%% in this paragraph. I can imagine roughly what the following
+ − 674
%%% sentence `The $\sigma$s' are given by the matchers for the
+ − 675
%%% $\alpha$s$ when matching $\varrho$s $\kappa$ against $\sigma$s
+ − 676
%%% $\kappa$.' means, but also think that it is too vague.
+ − 677
2443
+ − 678
\begin{isabelle}\ \ \ \ \ %%%
+ − 679
@{text "\<lambda>a b. map_prod (map_list a) b"}
+ − 680
\end{isabelle}
+ − 681
2233
+ − 682
\noindent
2274
+ − 683
which is essential in order to define the corresponding aggregate
+ − 684
abstraction and representation functions.
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 685
2247
+ − 686
To see how these definitions pan out in practise, let us return to our
+ − 687
example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
+ − 688
@{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
2274
+ − 689
fset"}. Feeding these types into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
2247
+ − 690
the abstraction function
2233
+ − 691
2455
+ − 692
\begin{isabelle}\ \ \ %%%
2444
+ − 693
\begin{tabular}{l}
+ − 694
@{text "(map_list (map_list id \<circ> Rep_fset) \<circ> Rep_fset) \<singlearr>"}\\
+ − 695
\mbox{}\hspace{4.5cm}@{text " Abs_fset \<circ> map_list id"}
+ − 696
\end{tabular}
2443
+ − 697
\end{isabelle}
2233
+ − 698
+ − 699
\noindent
2247
+ − 700
In our implementation we further
+ − 701
simplify this function by rewriting with the usual laws about @{text
2367
+ − 702
"map"}s and @{text "id"}, for example @{term "map_list id = id"} and @{text "f \<circ> id =
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 703
id \<circ> f = f"}. This gives us the simpler abstraction function
2237
+ − 704
2443
+ − 705
\begin{isabelle}\ \ \ \ \ %%%
+ − 706
@{text "(map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset"}
+ − 707
\end{isabelle}
2233
+ − 708
+ − 709
\noindent
+ − 710
which we can use for defining @{term "fconcat"} as follows
+ − 711
2443
+ − 712
\begin{isabelle}\ \ \ \ \ %%%
+ − 713
@{text "\<Union> \<equiv> ((map_list Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset) flat"}
+ − 714
\end{isabelle}
2232
+ − 715
2237
+ − 716
\noindent
2247
+ − 717
Note that by using the operator @{text "\<singlearr>"} and special clauses
+ − 718
for function types in \eqref{ABSREP}, we do not have to
2252
+ − 719
distinguish between arguments and results, but can deal with them uniformly.
+ − 720
Consequently, all definitions in the quotient package
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 721
are of the general form
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 722
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 723
\begin{isabelle}\ \ \ \ \ %%%
2554
+ − 724
\mbox{@{text "c \<equiv> ABS (\<sigma>, \<tau>) t"}}
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 725
\end{isabelle}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 726
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 727
\noindent
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 728
where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
2247
+ − 729
type of the defined quotient constant @{text "c"}. This data can be easily
+ − 730
generated from the declaration given by the user.
2252
+ − 731
To increase the confidence in this way of making definitions, we can prove
2247
+ − 732
that the terms involved are all typable.
2227
+ − 733
+ − 734
\begin{lemma}
+ − 735
If @{text "ABS (\<sigma>, \<tau>)"} returns some abstraction function @{text "Abs"}
+ − 736
and @{text "REP (\<sigma>, \<tau>)"} some representation function @{text "Rep"},
+ − 737
then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
+ − 738
@{text "\<tau> \<Rightarrow> \<sigma>"}.
+ − 739
\end{lemma}
2233
+ − 740
2237
+ − 741
\begin{proof}
2274
+ − 742
By mutual induction and analysing the definitions of @{text "ABS"} and @{text "REP"}.
+ − 743
The cases of equal types and function types are
2247
+ − 744
straightforward (the latter follows from @{text "\<singlearr>"} having the
+ − 745
type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
+ − 746
constructors we can observe that a map-function after applying the functions
+ − 747
@{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}. The
+ − 748
interesting case is the one with unequal type constructors. Since we know
+ − 749
the quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have
+ − 750
that @{text "Abs_\<kappa>\<^isub>q"} is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s
+ − 751
\<kappa>\<^isub>q"}. This type can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s
+ − 752
\<kappa>\<^isub>q"} where the type variables @{text "\<alpha>s"} are instantiated with the
+ − 753
@{text "\<tau>s"}. The complete type can be calculated by observing that @{text
+ − 754
"MAP (\<rho>s \<kappa>)"}, after applying the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it,
+ − 755
returns a term of type @{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is
+ − 756
equivalent to @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}, which we just have to compose with
+ − 757
@{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"} according to the type of @{text "\<circ>"}.\qed
2237
+ − 758
\end{proof}
2188
+ − 759
*}
+ − 760
2444
+ − 761
section {* Respectfulness and\\ Preservation \label{sec:resp} *}
2188
+ − 762
+ − 763
text {*
2445
+ − 764
\noindent
2247
+ − 765
The main point of the quotient package is to automatically ``lift'' theorems
+ − 766
involving constants over the raw type to theorems involving constants over
2274
+ − 767
the quotient type. Before we can describe this lifting process, we need to impose
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 768
two restrictions in form of proof obligations that arise during the
2274
+ − 769
lifting. The reason is that even if definitions for all raw constants
+ − 770
can be given, \emph{not} all theorems can be lifted to the quotient type. Most
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 771
notable is the bound variable function, that is the constant @{text bn},
2554
+ − 772
defined
2274
+ − 773
for raw lambda-terms as follows
2188
+ − 774
2444
+ − 775
\begin{isabelle}
+ − 776
\begin{center}
2252
+ − 777
@{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
2444
+ − 778
@{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\smallskip\\
2247
+ − 779
@{text "bn (\<lambda>x. t) \<equiv> {x} \<union> bn (t)"}
2444
+ − 780
\end{center}
2247
+ − 781
\end{isabelle}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 782
2247
+ − 783
\noindent
2274
+ − 784
We can generate a definition for this constant using @{text ABS} and @{text REP}.
+ − 785
But this constant does \emph{not} respect @{text "\<alpha>"}-equivalence and
2247
+ − 786
consequently no theorem involving this constant can be lifted to @{text
+ − 787
"\<alpha>"}-equated lambda terms. Homeier formulates the restrictions in terms of
2277
+ − 788
the properties of \emph{respectfulness} and \emph{preservation}. We have
2258
+ − 789
to slightly extend Homeier's definitions in order to deal with quotient
2247
+ − 790
compositions.
+ − 791
2422
+ − 792
%%% FIXME: Reviewer 3 asks why are the definitions that follow enough to deal
+ − 793
%%% with quotient composition.
+ − 794
2247
+ − 795
To formally define what respectfulness is, we have to first define
2419
+ − 796
the notion of \emph{aggregate equivalence relations} using the function @{text "REL(\<sigma>, \<tau>)"}
+ − 797
The idea behind this function is to simultaneously descend into the raw types
+ − 798
@{text \<sigma>} and quotient types @{text \<tau>}, and generate the appropriate
+ − 799
quotient equivalence relations in places where the types differ and equalities
+ − 800
elsewhere.
2274
+ − 801
+ − 802
\begin{center}
+ − 803
\hfill
2444
+ − 804
\begin{tabular}{l}
2554
+ − 805
\multicolumn{1}{@ {}l}{equal types:}
2444
+ − 806
@{text "REL (\<sigma>, \<sigma>)"} $\dn$ @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}\smallskip\\
+ − 807
\multicolumn{1}{@ {}l}{equal type constructors:}\\
+ − 808
@{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} $\dn$ @{text "rel_\<kappa> (REL (\<sigma>s, \<tau>s))"}\smallskip\\
+ − 809
\multicolumn{1}{@ {}l}{unequal type constructors:}\\
+ − 810
@{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} $\dn$ @{text "rel_\<kappa>\<^isub>q (REL (\<sigma>s', \<tau>s))"}\\
2274
+ − 811
\end{tabular}\hfill\numbered{REL}
+ − 812
\end{center}
2238
8ddf1330f2ed
completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 813
2274
+ − 814
\noindent
+ − 815
The @{text "\<sigma>s'"} in the last clause are calculated as in \eqref{ABSREP}:
2444
+ − 816
again we know that type @{text "\<alpha>s \<kappa>\<^isub>q"} is the quotient of the raw type
2367
+ − 817
@{text "\<rho>s \<kappa>"}. The @{text "\<sigma>s'"} are the substitutions for @{text "\<alpha>s"} obtained by matching
2274
+ − 818
@{text "\<rho>s \<kappa>"} and @{text "\<sigma>s \<kappa>"}.
+ − 819
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 820
Let us return to the lifting procedure of theorems. Assume we have a theorem
2274
+ − 821
that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
+ − 822
lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
+ − 823
constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation
2287
+ − 824
we generate the following proof obligation
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 825
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 826
\begin{isabelle}\ \ \ \ \ %%%
2554
+ − 827
@{text "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}.
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 828
\end{isabelle}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 829
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 830
\noindent
2277
+ − 831
Homeier calls these proof obligations \emph{respectfulness
2278
+ − 832
theorems}. However, unlike his quotient package, we might have several
+ − 833
respectfulness theorems for one constant---he has at most one.
+ − 834
The reason is that because of our quotient compositions, the types
2416
12283a96e851
The type does determine respectfulness, the constant without an instantiated type does not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 835
@{text \<sigma>} and @{text \<tau>} are not completely determined by @{text "c\<^bsub>r\<^esub>"}.
12283a96e851
The type does determine respectfulness, the constant without an instantiated type does not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 836
And for every instantiation of the types, a corresponding
12283a96e851
The type does determine respectfulness, the constant without an instantiated type does not.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 837
respectfulness theorem is necessary.
2278
+ − 838
+ − 839
Before lifting a theorem, we require the user to discharge
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 840
respectfulness proof obligations. In case of @{text bn}
2554
+ − 841
this obligation is %%as follows
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 842
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 843
\begin{isabelle}\ \ \ \ \ %%%
2443
+ − 844
@{text "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 845
\end{isabelle}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 846
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 847
\noindent
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 848
and the point is that the user cannot discharge it: because it is not true. To see this,
2274
+ − 849
we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun}
2412
+ − 850
using extensionality to obtain the false statement
2274
+ − 851
2443
+ − 852
\begin{isabelle}\ \ \ \ \ %%%
+ − 853
@{text "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 then bn(t\<^isub>1) = bn(t\<^isub>2)"}
+ − 854
\end{isabelle}
+ − 855
2274
+ − 856
\noindent
2554
+ − 857
In contrast, lifting a theorem about @{text "append"} to a theorem describing
+ − 858
the union of finite sets will mean to discharge the proof obligation
2274
+ − 859
2443
+ − 860
\begin{isabelle}\ \ \ \ \ %%%
+ − 861
@{text "(\<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub>) append append"}
+ − 862
\end{isabelle}
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 863
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 864
\noindent
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 865
To do so, we have to establish
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 866
2443
+ − 867
\begin{isabelle}\ \ \ \ \ %%%
2287
+ − 868
if @{text "xs \<approx>\<^bsub>list\<^esub> ys"} and @{text "us \<approx>\<^bsub>list\<^esub> vs"}
+ − 869
then @{text "xs @ us \<approx>\<^bsub>list\<^esub> ys @ vs"}
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 870
\end{isabelle}
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 871
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 872
\noindent
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 873
which is straightforward given the definition shown in \eqref{listequiv}.
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 874
2333
+ − 875
The second restriction we have to impose arises from non-lifted polymorphic
+ − 876
constants, which are instantiated to a type being quotient. For example,
+ − 877
take the @{term "cons"}-constructor to add a pair of natural numbers to a
+ − 878
list, whereby we assume the pair of natural numbers turns into an integer in
+ − 879
the quotient construction. The point is that we still want to use @{text
+ − 880
cons} for adding integers to lists---just with a different type. To be able
+ − 881
to lift such theorems, we need a \emph{preservation property} for @{text
+ − 882
cons}. Assuming we have a polymorphic raw constant @{text "c\<^isub>r :: \<sigma>"}
+ − 883
and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, then a
+ − 884
preservation property is as follows
+ − 885
2421
+ − 886
%%% FIXME: Reviewer 2 asks: You say what a preservation theorem is,
+ − 887
%%% but not which preservation theorems you assume. Do you generate a
+ − 888
%%% proof obligation for a preservation theorem for each raw constant
+ − 889
%%% and its corresponding lifted constant?
+ − 890
+ − 891
%%% Cezary: I think this would be a nice thing to do but we have not
+ − 892
%%% done it, the theorems need to be 'guessed' from the remaining obligations
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 893
2443
+ − 894
\begin{isabelle}\ \ \ \ \ %%%
2455
+ − 895
@{text "Quot R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
2443
+ − 896
\end{isabelle}
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 897
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 898
\noindent
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 899
where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
2278
+ − 900
In case of @{text cons} (which has type @{text "\<alpha> \<Rightarrow> \<alpha> list \<Rightarrow> \<alpha> list"}) we have
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 901
2443
+ − 902
\begin{isabelle}\ \ \ \ \ %%%
2445
+ − 903
@{text "(Rep \<singlearr> map_list Rep \<singlearr> map_list Abs) cons = cons"}
2443
+ − 904
\end{isabelle}
2274
+ − 905
+ − 906
\noindent
2455
+ − 907
under the assumption @{term "Quotient R Abs Rep"}. The point is that if we have
2278
+ − 908
an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
+ − 909
with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
2445
+ − 910
then we need to show this preservation property.
2274
+ − 911
2279
+ − 912
%%%@ {thm [display, indent=10] insert_preserve2[no_vars]}
2189
+ − 913
2278
+ − 914
%Given two quotients, one of which quotients a container, and the
+ − 915
%other quotients the type in the container, we can write the
+ − 916
%composition of those quotients. To compose two quotient theorems
+ − 917
%we compose the relations with relation composition as defined above
+ − 918
%and the abstraction and relation functions are the ones of the sub
+ − 919
%quotients composed with the usual function composition.
+ − 920
%The @ {term "Rep"} and @ {term "Abs"} functions that we obtain agree
+ − 921
%with the definition of aggregate Abs/Rep functions and the
+ − 922
%relation is the same as the one given by aggregate relations.
+ − 923
%This becomes especially interesting
+ − 924
%when we compose the quotient with itself, as there is no simple
+ − 925
%intermediate step.
+ − 926
%
+ − 927
%Lets take again the example of @ {term flat}. To be able to lift
+ − 928
%theorems that talk about it we provide the composition quotient
+ − 929
%theorem which allows quotienting inside the container:
+ − 930
%
+ − 931
%If @ {term R} is an equivalence relation and @ {term "Quotient R Abs Rep"}
+ − 932
%then
+ − 933
%
2367
+ − 934
%@ {text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map_list Abs) (map_list Rep o rep_fset)"}
2275
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 935
%%%
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 936
%%%\noindent
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 937
%%%this theorem will then instantiate the quotients needed in the
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 938
%%%injection and cleaning proofs allowing the lifting procedure to
69b80ad616c5
finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 939
%%%proceed in an unchanged way.
2192
+ − 940
*}
+ − 941
2256
+ − 942
section {* Lifting of Theorems\label{sec:lift} *}
1978
+ − 943
2194
+ − 944
text {*
2271
+ − 945
2422
+ − 946
%%% FIXME Reviewer 3 asks: Section 5 shows the technicalities of
+ − 947
%%% lifting theorems. But there is no clarification about the
+ − 948
%%% correctness. A reader would also be interested in seeing some
+ − 949
%%% discussions about the generality and limitation of the approach
+ − 950
%%% proposed there
+ − 951
2445
+ − 952
\noindent
2278
+ − 953
The main benefit of a quotient package is to lift automatically theorems over raw
+ − 954
types to theorems over quotient types. We will perform this lifting in
+ − 955
three phases, called \emph{regularization},
2287
+ − 956
\emph{injection} and \emph{cleaning} according to procedures in Homeier's ML-code.
2445
+ − 957
Space restrictions, unfortunately, prevent us from giving anything but a sketch of these three
+ − 958
phases. However, we will precisely define the input and output data of these phases
+ − 959
(this was omitted in \cite{Homeier05}).
2271
+ − 960
2278
+ − 961
The purpose of regularization is to change the quantifiers and abstractions
2412
+ − 962
in a ``raw'' theorem to quantifiers over variables that respect their respective relations
2279
+ − 963
(Definition \ref{def:respects} states what respects means). The purpose of injection is to add @{term Rep}
2271
+ − 964
and @{term Abs} of appropriate types in front of constants and variables
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 965
of the raw type so that they can be replaced by the corresponding constants from the
2279
+ − 966
quotient type. The purpose of cleaning is to bring the theorem derived in the
+ − 967
first two phases into the form the user has specified. Abstractly, our
+ − 968
package establishes the following three proof steps:
+ − 969
2421
+ − 970
%%% FIXME: Reviewer 1 complains that the reader needs to guess the
+ − 971
%%% meaning of reg_thm and inj_thm, as well as the arguments of REG
+ − 972
%%% which are given above. I wouldn't change it.
+ − 973
2279
+ − 974
\begin{center}
2369
+ − 975
\begin{tabular}{l@ {\hspace{4mm}}l}
2371
+ − 976
1.) Regularization & @{text "raw_thm \<longrightarrow> reg_thm"}\\
2369
+ − 977
2.) Injection & @{text "reg_thm \<longleftrightarrow> inj_thm"}\\
+ − 978
3.) Cleaning & @{text "inj_thm \<longleftrightarrow> quot_thm"}\\
2279
+ − 979
\end{tabular}
+ − 980
\end{center}
2193
+ − 981
2279
+ − 982
\noindent
2371
+ − 983
which means, stringed together, the raw theorem implies the quotient theorem.
+ − 984
In contrast to other quotient packages, our package requires that the user specifies
+ − 985
both, the @{text "raw_thm"} (as theorem) and the \emph{term} of the @{text "quot_thm"}.\footnote{Though we
2279
+ − 986
also provide a fully automated mode, where the @{text "quot_thm"} is guessed
2371
+ − 987
from the form of @{text "raw_thm"}.} As a result, the user has fine control
+ − 988
over which parts of a raw theorem should be lifted.
2279
+ − 989
2371
+ − 990
The second and third proof step performed in package will always succeed if the appropriate
2279
+ − 991
respectfulness and preservation theorems are given. In contrast, the first
+ − 992
proof step can fail: a theorem given by the user does not always
2319
+ − 993
imply a regularized version and a stronger one needs to be proved. An example
2287
+ − 994
for this kind of failure is the simple statement for integers @{text "0 \<noteq> 1"}.
+ − 995
One might hope that it can be proved by lifting @{text "(0, 0) \<noteq> (1, 0)"},
2371
+ − 996
but this raw theorem only shows that two particular elements in the
2319
+ − 997
equivalence classes are not equal. In order to obtain @{text "0 \<noteq> 1"}, a
+ − 998
more general statement stipulating that the equivalence classes are not
+ − 999
equal is necessary. This kind of failure is beyond the scope where the
+ − 1000
quotient package can help: the user has to provide a raw theorem that
+ − 1001
can be regularized automatically, or has to provide an explicit proof
2445
+ − 1002
for the first proof step. Homeier gives more details about this issue
+ − 1003
in the long version of \cite{Homeier05}.
2193
+ − 1004
2271
+ − 1005
In the following we will first define the statement of the
2280
229660b9f2fc
Fix regularization. Two "FIXME" left in introduction. Minor spellings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1006
regularized theorem based on @{text "raw_thm"} and
2279
+ − 1007
@{text "quot_thm"}. Then we define the statement of the injected theorem, based
2287
+ − 1008
on @{text "reg_thm"} and @{text "quot_thm"}. We then show the three proof steps,
2279
+ − 1009
which can all be performed independently from each other.
2197
+ − 1010
2373
+ − 1011
We first define the function @{text REG}, which takes the terms of the
2371
+ − 1012
@{text "raw_thm"} and @{text "quot_thm"} as input and returns
2373
+ − 1013
@{text "reg_thm"}. The idea
2207
+ − 1014
behind this function is that it replaces quantifiers and
+ − 1015
abstractions involving raw types by bounded ones, and equalities
2373
+ − 1016
involving raw types by appropriate aggregate
2445
+ − 1017
equivalence relations. It is defined by simultaneous recursion on
+ − 1018
the structure of the terms of @{text "raw_thm"} and @{text "quot_thm"} as follows:
+ − 1019
%
2244
+ − 1020
\begin{center}
2445
+ − 1021
\begin{tabular}{@ {}l@ {}}
2554
+ − 1022
\multicolumn{1}{@ {}l@ {}}{abstractions:}\\%%\smallskip\\
2444
+ − 1023
@{text "REG (\<lambda>x\<^sup>\<sigma>. t, \<lambda>x\<^sup>\<tau>. s)"} $\dn$
2273
+ − 1024
$\begin{cases}
+ − 1025
@{text "\<lambda>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
2445
+ − 1026
@{text "\<lambda>x\<^sup>\<sigma> \<in> Resp (REL (\<sigma>, \<tau>)). REG (t, s)"}
2554
+ − 1027
\end{cases}$\\%%\smallskip\\
2445
+ − 1028
\multicolumn{1}{@ {}l@ {}}{universal quantifiers:}\\
2444
+ − 1029
@{text "REG (\<forall>x\<^sup>\<sigma>. t, \<forall>x\<^sup>\<tau>. s)"} $\dn$
2273
+ − 1030
$\begin{cases}
+ − 1031
@{text "\<forall>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
2445
+ − 1032
@{text "\<forall>x\<^sup>\<sigma> \<in> Resp (REL (\<sigma>, \<tau>)). REG (t, s)"}
2554
+ − 1033
\end{cases}$\\%%\smallskip\\
+ − 1034
\multicolumn{1}{@ {}l@ {}}{equality: \hspace{3mm}%%}\smallskip\\
2280
229660b9f2fc
Fix regularization. Two "FIXME" left in introduction. Minor spellings.
Cezary Kaliszyk <kaliszyk@in.tum.de>
diff
changeset
+ − 1035
%% REL of two equal types is the equality so we do not need a separate case
2554
+ − 1036
@{text "REG (=\<^bsup>\<sigma>\<Rightarrow>\<sigma>\<Rightarrow>bool\<^esup>, =\<^bsup>\<tau>\<Rightarrow>\<tau>\<Rightarrow>bool\<^esup>)"} $\dn$ @{text "REL (\<sigma>, \<tau>)"}}\smallskip\\
2445
+ − 1037
\multicolumn{1}{@ {}l@ {}}{applications, variables and constants:}\\
2444
+ − 1038
@{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)"}\\
+ − 1039
@{text "REG (x\<^isub>1, x\<^isub>2)"} $\dn$ @{text "x\<^isub>1"}\\
+ − 1040
@{text "REG (c\<^isub>1, c\<^isub>2)"} $\dn$ @{text "c\<^isub>1"}\\
2371
+ − 1041
\end{tabular}
2244
+ − 1042
\end{center}
2273
+ − 1043
%
+ − 1044
\noindent
2230
+ − 1045
In the above definition we omitted the cases for existential quantifiers
2207
+ − 1046
and unique existential quantifiers, as they are very similar to the cases
2376
+ − 1047
for the universal quantifier.
2279
+ − 1048
+ − 1049
Next we define the function @{text INJ} which takes as argument
+ − 1050
@{text "reg_thm"} and @{text "quot_thm"} (both as
+ − 1051
terms) and returns @{text "inj_thm"}:
2198
+ − 1052
2245
+ − 1053
\begin{center}
2444
+ − 1054
\begin{tabular}{l}
+ − 1055
\multicolumn{1}{@ {}l}{abstractions:}\\
+ − 1056
@{text "INJ (\<lambda>x. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} $\dn$\\
+ − 1057
\hspace{18mm}$\begin{cases}
2273
+ − 1058
@{text "\<lambda>x. INJ (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
+ − 1059
@{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x. INJ (t, s)))"}
2444
+ − 1060
\end{cases}$\smallskip\\
+ − 1061
@{text "INJ (\<lambda>x \<in> R. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} $\dn$\\
+ − 1062
\hspace{18mm}@{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x \<in> R. INJ (t, s)))"}\smallskip\\
+ − 1063
\multicolumn{1}{@ {}l}{universal quantifiers:}\\
+ − 1064
@{text "INJ (\<forall> t, \<forall> s) "} $\dn$ @{text "\<forall> INJ (t, s)"}\\
+ − 1065
@{text "INJ (\<forall> t \<in> R, \<forall> s) "} $\dn$ @{text "\<forall> INJ (t, s) \<in> R"}\smallskip\\
+ − 1066
\multicolumn{1}{@ {}l}{applications, variables and constants:}\smallskip\\
+ − 1067
@{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)"}\\
+ − 1068
@{text "INJ (x\<^isub>1\<^sup>\<sigma>, x\<^isub>2\<^sup>\<tau>) "} $\dn$
2273
+ − 1069
$\begin{cases}
+ − 1070
@{text "x\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
+ − 1071
@{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) x\<^isub>1)"}\\
+ − 1072
\end{cases}$\\
2444
+ − 1073
@{text "INJ (c\<^isub>1\<^sup>\<sigma>, c\<^isub>2\<^sup>\<tau>) "} $\dn$
2273
+ − 1074
$\begin{cases}
+ − 1075
@{text "c\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
+ − 1076
@{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) c\<^isub>1)"}\\
+ − 1077
\end{cases}$\\
2245
+ − 1078
\end{tabular}
+ − 1079
\end{center}
2198
+ − 1080
2279
+ − 1081
\noindent
2319
+ − 1082
In this definition we again omitted the cases for existential and unique existential
2376
+ − 1083
quantifiers.
2208
+ − 1084
2421
+ − 1085
%%% FIXME: Reviewer2 citing following sentence: You mention earlier
+ − 1086
%%% that this implication may fail to be true. Does that meant that
+ − 1087
%%% the `first proof step' is a heuristic that proves the implication
+ − 1088
%%% raw_thm \implies reg_thm in some instances, but fails in others?
+ − 1089
%%% You should clarify under which circumstances the implication is
+ − 1090
%%% being proved here.
2423
+ − 1091
%%% Cezary: It would be nice to cite Homeiers discussions in the
+ − 1092
%%% Quotient Package manual from HOL (the longer paper), do you agree?
2421
+ − 1093
2445
+ − 1094
In the first phase, establishing @{text "raw_thm \<longrightarrow> reg_thm"}, we always
2279
+ − 1095
start with an implication. Isabelle provides \emph{mono} rules that can split up
2319
+ − 1096
the implications into simpler implicational subgoals. This succeeds for every
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1097
monotone connective, except in places where the function @{text REG} replaced,
2445
+ − 1098
for instance, a quantifier by a bounded quantifier. To decompose them, we have
+ − 1099
to prove that the relations involved are aggregate equivalence relations.
2208
+ − 1100
2445
+ − 1101
+ − 1102
%In this case we have
+ − 1103
%rules of the form
+ − 1104
%
+ − 1105
% \begin{isabelle}\ \ \ \ \ %%%
+ − 1106
%@{text "(\<forall>x. R x \<longrightarrow> (P x \<longrightarrow> Q x)) \<longrightarrow> (\<forall>x. P x \<longrightarrow> \<forall>x \<in> R. Q x)"}
+ − 1107
%\end{isabelle}
2279
+ − 1108
2445
+ − 1109
%\noindent
+ − 1110
%They decompose a bounded quantifier on the right-hand side. We can decompose a
+ − 1111
%bounded quantifier anywhere if R is an equivalence relation or
+ − 1112
%if it is a relation over function types with the range being an equivalence
+ − 1113
%relation. If @{text R} is an equivalence relation we can prove that
2261
+ − 1114
2445
+ − 1115
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1116
%@{text "\<forall>x \<in> Resp R. P x = \<forall>x. P x"}
+ − 1117
%\end{isabelle}
2199
+ − 1118
2445
+ − 1119
%\noindent
+ − 1120
%If @{term R\<^isub>2} is an equivalence relation, we can prove that for any predicate @{term P}
2231
+ − 1121
2421
+ − 1122
%%% FIXME Reviewer 1 claims the theorem is obviously false so maybe we
+ − 1123
%%% should include a proof sketch?
+ − 1124
2445
+ − 1125
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1126
%@{thm (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
+ − 1127
%\end{isabelle}
2209
+ − 1128
2445
+ − 1129
%\noindent
+ − 1130
%The last theorem is new in comparison with Homeier's package. There the
+ − 1131
%injection procedure would be used to prove such goals and
+ − 1132
%the assumption about the equivalence relation would be used. We use the above theorem directly,
+ − 1133
%because this allows us to completely separate the first and the second
+ − 1134
%proof step into two independent ``units''.
2206
+ − 1135
2445
+ − 1136
The second phase, establishing @{text "reg_thm \<longleftrightarrow> inj_thm"}, starts with an equality
2412
+ − 1137
between the terms of the regularized theorem and the injected theorem.
2279
+ − 1138
The proof again follows the structure of the
2445
+ − 1139
two underlying terms taking respectfulness theorems into account.
2199
+ − 1140
2445
+ − 1141
%\begin{itemize}
+ − 1142
%\item For two constants an appropriate respectfulness theorem is applied.
+ − 1143
%\item For two variables, we use the assumptions proved in the regularization step.
+ − 1144
%\item For two abstractions, we @{text "\<eta>"}-expand and @{text "\<beta>"}-reduce them.
+ − 1145
%\item For two applications, we check that the right-hand side is an application of
+ − 1146
% @{term Rep} to an @{term Abs} and @{term "Quotient R Rep Abs"} holds. If yes then we
+ − 1147
% can apply the theorem:
2271
+ − 1148
2445
+ − 1149
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1150
% @{term "R x y \<longrightarrow> R x (Rep (Abs y))"}
+ − 1151
%\end{isabelle}
2271
+ − 1152
2445
+ − 1153
% Otherwise we introduce an appropriate relation between the subterms
+ − 1154
% and continue with two subgoals using the lemma:
2271
+ − 1155
2445
+ − 1156
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1157
% @{text "(R\<^isub>1 \<doublearr> R\<^isub>2) f g \<longrightarrow> R\<^isub>1 x y \<longrightarrow> R\<^isub>2 (f x) (g y)"}
+ − 1158
%\end{isabelle}
+ − 1159
%\end{itemize}
2199
+ − 1160
2279
+ − 1161
We defined the theorem @{text "inj_thm"} in such a way that
2445
+ − 1162
establishing in the third phase the equivalence @{text "inj_thm \<longleftrightarrow> quot_thm"} can be
2279
+ − 1163
achieved by rewriting @{text "inj_thm"} with the preservation theorems and quotient
2445
+ − 1164
definitions. This step also requires that the definitions of all lifted constants
+ − 1165
are used to fold the @{term Rep} with the raw constants. We will give more details
+ − 1166
about our lifting procedure in a longer version of this paper.
+ − 1167
+ − 1168
%Next for
+ − 1169
%all abstractions and quantifiers the lambda and
+ − 1170
%quantifier preservation theorems are used to replace the
+ − 1171
%variables that include raw types with respects by quantifiers
+ − 1172
%over variables that include quotient types. We show here only
+ − 1173
%the lambda preservation theorem. Given
+ − 1174
%@{term "Quotient R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1"} and @{term "Quotient R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2"}, we have:
2211
+ − 1175
2445
+ − 1176
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1177
%@{thm (concl) lambda_prs[of _ "Abs\<^isub>1" "Rep\<^isub>1" _ "Abs\<^isub>2" "Rep\<^isub>2", no_vars]}
+ − 1178
%\end{isabelle}
2199
+ − 1179
2445
+ − 1180
%\noindent
+ − 1181
%Next, relations over lifted types can be rewritten to equalities
+ − 1182
%over lifted type. Rewriting is performed with the following theorem,
+ − 1183
%which has been shown by Homeier~\cite{Homeier05}:
2211
+ − 1184
2445
+ − 1185
%\begin{isabelle}\ \ \ \ \ %%%
+ − 1186
%@{thm (concl) Quotient_rel_rep[no_vars]}
+ − 1187
%\end{isabelle}
2199
+ − 1188
2445
+ − 1189
+ − 1190
%Finally, we rewrite with the preservation theorems. This will result
+ − 1191
%in two equal terms that can be solved by reflexivity.
+ − 1192
*}
1994
+ − 1193
2553
+ − 1194
section {* Conclusion and Related Work\label{sec:conc}*}
2376
+ − 1195
2553
+ − 1196
text {*
+ − 1197
+ − 1198
\noindent
+ − 1199
The code of the quotient package and the examples described here are already
+ − 1200
included in the standard distribution of Isabelle.
+ − 1201
\footnote{Available from \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.}
+ − 1202
The package is
+ − 1203
heavily used in the new version of Nominal Isabelle, which provides a
+ − 1204
convenient reasoning infrastructure for programming language calculi
+ − 1205
involving general binders. To achieve this, it builds types representing
+ − 1206
@{text \<alpha>}-equivalent terms. Earlier versions of Nominal Isabelle have been
+ − 1207
used successfully in formalisations of an equivalence checking algorithm for
+ − 1208
LF \cite{UrbanCheneyBerghofer08}, Typed
+ − 1209
Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for concurrency
+ − 1210
\cite{BengtsonParow09} and a strong normalisation result for cut-elimination
+ − 1211
in classical logic \cite{UrbanZhu08}.
+ − 1212
+ − 1213
+ − 1214
There is a wide range of existing literature for dealing with quotients
+ − 1215
in theorem provers. Slotosch~\cite{Slotosch97} implemented a mechanism that
+ − 1216
automatically defines quotient types for Isabelle/HOL. But he did not
+ − 1217
include theorem lifting. Harrison's quotient package~\cite{harrison-thesis}
+ − 1218
is the first one that is able to automatically lift theorems, however only
+ − 1219
first-order theorems (that is theorems where abstractions, quantifiers and
+ − 1220
variables do not involve functions that include the quotient type). There is
+ − 1221
also some work on quotient types in non-HOL based systems and logical
+ − 1222
frameworks, including theory interpretations in
+ − 1223
PVS~\cite{PVS:Interpretations}, new types in MetaPRL~\cite{Nogin02}, and
+ − 1224
setoids in Coq \cite{ChicliPS02}. Paulson showed a construction of
+ − 1225
quotients that does not require the Hilbert Choice operator, but also only
+ − 1226
first-order theorems can be lifted~\cite{Paulson06}. The most related work
+ − 1227
to our package is the package for HOL4 by Homeier~\cite{Homeier05}. He
+ − 1228
introduced most of the abstract notions about quotients and also deals with
+ − 1229
lifting of higher-order theorems. However, he cannot deal with quotient
+ − 1230
compositions (needed for lifting theorems about @{text flat}). Also, a
+ − 1231
number of his definitions, like @{text ABS}, @{text REP} and @{text INJ} etc
+ − 1232
only exist in \cite{Homeier05} as ML-code, not included in the paper.
+ − 1233
Like Homeier's, our quotient package can deal with partial equivalence
+ − 1234
relations, but for lack of space we do not describe the mechanisms
+ − 1235
needed for this kind of quotient constructions.
+ − 1236
+ − 1237
%%% FIXME Reviewer 3 would like to know more about the lifting in Coq and PVS,
+ − 1238
%%% and some comparison. I don't think we have the space for any additions...
+ − 1239
+ − 1240
One feature of our quotient package is that when lifting theorems, the user
+ − 1241
can precisely specify what the lifted theorem should look like. This feature
+ − 1242
is necessary, for example, when lifting an induction principle for two
+ − 1243
lists. Assuming this principle has as the conclusion a predicate of the
+ − 1244
form @{text "P xs ys"}, then we can precisely specify whether we want to
+ − 1245
quotient @{text "xs"} or @{text "ys"}, or both. We found this feature very
+ − 1246
useful in the new version of Nominal Isabelle, where such a choice is
+ − 1247
required to generate a reasoning infrastructure for alpha-equated terms.
+ − 1248
%%
+ − 1249
%% give an example for this
+ − 1250
%%
2554
+ − 1251
\smallskip
2553
+ − 1252
+ − 1253
\noindent
+ − 1254
{\bf Acknowledgements:} We would like to thank Peter Homeier for the many
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1255
discussions about his HOL4 quotient package and explaining to us
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1256
some of its finer points in the implementation. Without his patient
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1257
help, this work would have been impossible.
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1258
*}
2553
+ − 1259
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1260
text_raw {*
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1261
%%\bibliographystyle{abbrv}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1262
\bibliography{root}
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1263
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1264
\appendix
2553
+ − 1265
+ − 1266
*}
+ − 1267
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1268
section {* Examples \label{sec:examples} *}
1994
+ − 1269
2421
+ − 1270
text {*
2206
+ − 1271
2421
+ − 1272
%%% FIXME Reviewer 1 would like an example of regularized and injected
+ − 1273
%%% statements. He asks for the examples twice, but I would still ignore
+ − 1274
%%% it due to lack of space...
2210
+ − 1275
2445
+ − 1276
\noindent
2553
+ − 1277
In this appendix we will show a sequence of declarations for defining the
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1278
type of integers by quotienting pairs of natural numbers, and
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1279
lifting one theorem.
2210
+ − 1280
2240
+ − 1281
A user of our quotient package first needs to define a relation on
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1282
the raw type with which the quotienting will be performed. We give
2279
+ − 1283
the same integer relation as the one presented in \eqref{natpairequiv}:
2240
+ − 1284
2443
+ − 1285
\begin{isabelle}\ \ \ \ \ %
2279
+ − 1286
\begin{tabular}{@ {}l}
+ − 1287
\isacommand{fun}~~@{text "int_rel :: (nat \<times> nat) \<Rightarrow> (nat \<times> nat) \<Rightarrow> (nat \<times> nat)"}\\
+ − 1288
\isacommand{where}~~@{text "int_rel (m, n) (p, q) = (m + q = n + p)"}
+ − 1289
\end{tabular}
2239
+ − 1290
\end{isabelle}
2210
+ − 1291
2239
+ − 1292
\noindent
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1293
Next the quotient type must be defined. This generates a proof obligation that the
2279
+ − 1294
relation is an equivalence relation, which is solved automatically using the
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1295
definition of equivalence and extensionality:
2210
+ − 1296
2443
+ − 1297
\begin{isabelle}\ \ \ \ \ %
2279
+ − 1298
\begin{tabular}{@ {}l}
+ − 1299
\isacommand{quotient\_type}~~@{text "int"}~~\isacommand{=}~~@{text "(nat \<times> nat)"}~~\isacommand{/}~~@{text "int_rel"}\\
+ − 1300
\hspace{5mm}@{text "by (auto simp add: equivp_def expand_fun_eq)"}
+ − 1301
\end{tabular}
2239
+ − 1302
\end{isabelle}
2210
+ − 1303
2239
+ − 1304
\noindent
2210
+ − 1305
The user can then specify the constants on the quotient type:
+ − 1306
2443
+ − 1307
\begin{isabelle}\ \ \ \ \ %
2240
+ − 1308
\begin{tabular}{@ {}l}
2279
+ − 1309
\isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\[3mm]
2443
+ − 1310
\isacommand{fun}~~@{text "add_pair"}\\
+ − 1311
\isacommand{where}~~%
2287
+ − 1312
@{text "add_pair (m, n) (p, q) \<equiv> (m + p :: nat, n + q :: nat)"}\\
2273
+ − 1313
\isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
2287
+ − 1314
\isacommand{is}~~@{text "add_pair"}\\
2240
+ − 1315
\end{tabular}
+ − 1316
\end{isabelle}
2210
+ − 1317
2240
+ − 1318
\noindent
2279
+ − 1319
The following theorem about addition on the raw level can be proved.
2210
+ − 1320
2443
+ − 1321
\begin{isabelle}\ \ \ \ \ %
2287
+ − 1322
\isacommand{lemma}~~@{text "add_pair_zero: int_rel (add_pair (0, 0) x) x"}
2240
+ − 1323
\end{isabelle}
2210
+ − 1324
2240
+ − 1325
\noindent
2413
+ − 1326
If the user lifts this theorem, the quotient package performs all the lifting
+ − 1327
automatically leaving the respectfulness proof for the constant @{text "add_pair"}
+ − 1328
as the only remaining proof obligation. This property needs to be proved by the user:
2210
+ − 1329
2443
+ − 1330
\begin{isabelle}\ \ \ \ \ %
2279
+ − 1331
\begin{tabular}{@ {}l}
2287
+ − 1332
\isacommand{lemma}~~@{text "[quot_respect]:"}\\
+ − 1333
@{text "(int_rel \<doublearr> int_rel \<doublearr> int_rel) add_pair add_pair"}
2279
+ − 1334
\end{tabular}
2240
+ − 1335
\end{isabelle}
+ − 1336
+ − 1337
\noindent
2413
+ − 1338
It can be discharged automatically by Isabelle when hinting to unfold the definition
2273
+ − 1339
of @{text "\<doublearr>"}.
2332
9a560e489c64
polished paper again (and took out some claims about Homeier's package)
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1340
After this, the user can prove the lifted lemma as follows:
2210
+ − 1341
2443
+ − 1342
\begin{isabelle}\ \ \ \ \ %
2287
+ − 1343
\isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting add_pair_zero"}
2240
+ − 1344
\end{isabelle}
2210
+ − 1345
2240
+ − 1346
\noindent
2333
+ − 1347
or by using the completely automated mode stating just:
2210
+ − 1348
2443
+ − 1349
\begin{isabelle}\ \ \ \ \ %
2287
+ − 1350
\isacommand{thm}~~@{text "add_pair_zero[quot_lifted]"}
2240
+ − 1351
\end{isabelle}
2210
+ − 1352
2240
+ − 1353
\noindent
2443
+ − 1354
Both methods give the same result, namely @{text "0 + x = x"}
2333
+ − 1355
where @{text x} is of type integer.
2279
+ − 1356
Although seemingly simple, arriving at this result without the help of a quotient
2333
+ − 1357
package requires a substantial reasoning effort (see \cite{Paulson06}).
2210
+ − 1358
*}
2558
6cfb5d8a5b5b
expanded the paper by uncommenting the comments and adding the appendix
Christian Urban <urbanc@in.tum.de>
diff
changeset
+ − 1359
2224
+ − 1360
2227
+ − 1361
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1362
(*<*)
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
+ − 1363
end
1978
+ − 1364
(*>*)