author | Christian Urban <urbanc@in.tum.de> |
Tue, 11 May 2010 12:18:26 +0100 | |
changeset 2102 | 200954544cae |
parent 1994 | abada9e6f943 |
child 2103 | e08e3c29dbc0 |
permissions | -rw-r--r-- |
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 |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
imports "Quotient" |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
"LaTeXsugar" |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
begin |
1994 | 6 |
|
7 |
notation (latex output) |
|
8 |
fun_rel ("_ ===> _" [51, 51] 50) |
|
9 |
||
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
(*>*) |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
11 |
|
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
section {* Introduction *} |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
|
2102
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
14 |
text {* |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
15 |
Isabelle is a generic theorem prover in which many logics can be implemented. |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
16 |
The most widely used one, however, is |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
17 |
Higher-Order Logic (HOL). This logic consists of a small number of |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
18 |
axioms and inference |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
19 |
rules over a simply-typed term-language. Safe reasoning in HOL is ensured by two very restricted |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
20 |
mechanisms for extending the logic: one is the definition of new constants |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
21 |
in terms of existing ones; the other is the introduction of new types |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
22 |
by identifying non-empty subsets in existing types. It is well understood |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
23 |
to use both mechanism for dealing with quotient constructions in HOL (cite Larry). |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
24 |
For example the integers in Isabelle/HOL are constructed by a quotient construction over |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
25 |
the type @{typ "nat \<times> nat"} and the equivalence relation |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
26 |
|
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
27 |
@{text [display] "(n\<^isub>1, n\<^isub>2) \<approx> (m\<^isub>1, m\<^isub>2) \<equiv> n\<^isub>1 - n \<^isub>2 = m\<^isub>1 - m \<^isub>2"} |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
28 |
|
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
29 |
\noindent |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
30 |
The problem is that one |
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
31 |
|
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
32 |
|
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
33 |
|
200954544cae
added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents:
1994
diff
changeset
|
34 |
*} |
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
|
1978 | 36 |
subsection {* Contributions *} |
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
|
1978 | 38 |
text {* |
39 |
We present the detailed lifting procedure, which was not shown before. |
|
40 |
||
41 |
The quotient package presented in this paper has the following |
|
42 |
advantages over existing packages: |
|
43 |
\begin{itemize} |
|
44 |
||
45 |
\item We define quotient composition, function map composition and |
|
46 |
relation map composition. This lets lifting polymorphic types with |
|
47 |
subtypes quotiented as well. We extend the notions of |
|
48 |
respectfullness and preservation to cope with quotient |
|
49 |
composition. |
|
50 |
||
51 |
\item We allow lifting only some occurrences of quotiented |
|
52 |
types. Rsp/Prs extended. (used in nominal) |
|
53 |
||
54 |
\item We regularize more thanks to new lemmas. (inductions in |
|
55 |
nominal). |
|
56 |
||
57 |
\item The quotient package is very modular. Definitions can be added |
|
58 |
separately, rsp and prs can be proved separately and theorems can |
|
59 |
be lifted on a need basis. (useful with type-classes). |
|
60 |
||
61 |
\item Can be used both manually (attribute, separate tactics, |
|
62 |
rsp/prs databases) and programatically (automated definition of |
|
63 |
lifted constants, the rsp proof obligations and theorem statement |
|
64 |
translation according to given quotients). |
|
65 |
||
66 |
\end{itemize} |
|
67 |
*} |
|
68 |
||
69 |
section {* Quotient Type*} |
|
70 |
||
71 |
text {* |
|
72 |
Defintion of quotient, |
|
73 |
||
74 |
Equivalence, |
|
75 |
||
76 |
Relation map and function map |
|
77 |
*} |
|
78 |
||
79 |
section {* Constants *} |
|
80 |
||
81 |
text {* |
|
82 |
Rep and Abs, Rsp and Prs |
|
83 |
*} |
|
84 |
||
85 |
section {* Lifting Theorems *} |
|
86 |
||
1994 | 87 |
text {* TBD *} |
88 |
||
89 |
text {* Why providing a statement to prove is necessary is some cases *} |
|
90 |
||
91 |
subsection {* Regularization *} |
|
92 |
||
93 |
text {* |
|
94 |
Transformation of the theorem statement: |
|
95 |
\begin{itemize} |
|
96 |
\item Quantifiers and abstractions involving raw types replaced by bounded ones. |
|
97 |
\item Equalities involving raw types replaced by bounded ones. |
|
98 |
\end{itemize} |
|
99 |
||
100 |
The procedure. |
|
101 |
||
102 |
Example of non-regularizable theorem ($0 = 1$). |
|
103 |
||
104 |
New regularization lemmas: |
|
105 |
\begin{lemma} |
|
106 |
If @{term R2} is an equivalence relation, then: |
|
107 |
\begin{eqnarray} |
|
108 |
@{thm (rhs) ball_reg_eqv_range[no_vars]} & = & @{thm (lhs) ball_reg_eqv_range[no_vars]}\\ |
|
109 |
@{thm (rhs) bex_reg_eqv_range[no_vars]} & = & @{thm (lhs) bex_reg_eqv_range[no_vars]} |
|
110 |
\end{eqnarray} |
|
111 |
\end{lemma} |
|
112 |
||
113 |
*} |
|
114 |
||
115 |
subsection {* Injection *} |
|
116 |
||
117 |
subsection {* Cleaning *} |
|
118 |
||
119 |
text {* Preservation of quantifiers, abstractions, relations, quotient-constants |
|
120 |
(definitions) and user given constant preservation lemmas *} |
|
121 |
||
122 |
section {* Examples *} |
|
123 |
||
1978 | 124 |
section {* Related Work *} |
125 |
||
126 |
text {* |
|
127 |
\begin{itemize} |
|
128 |
||
129 |
\item Peter Homeier's package (and related work from there), John |
|
130 |
Harrison's one. |
|
131 |
||
132 |
\item Manually defined quotients in Isabelle/HOL Library (Larry's |
|
133 |
quotients, Markus's Quotient\_Type, Dixon's FSet, \ldots) |
|
134 |
||
135 |
\item Oscar Slotosch defines quotient-type automatically but no |
|
136 |
lifting. |
|
137 |
||
138 |
\item PER. And how to avoid it. |
|
139 |
||
140 |
\item Necessity of Hilbert Choice op. |
|
141 |
||
142 |
\item Setoids in Coq |
|
143 |
||
144 |
\end{itemize} |
|
145 |
*} |
|
1975
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
146 |
|
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
147 |
(*<*) |
b1281a0051ae
added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
148 |
end |
1978 | 149 |
(*>*) |