Quotient-Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Fri, 14 May 2010 17:58:26 +0100
changeset 2137 5b007ac41b29
parent 2103 e08e3c29dbc0
child 2152 d7d4491535a9
permissions -rw-r--r--
moved old parser and fv into attic
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
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
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     6
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     7
notation (latex output)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     8
  fun_rel ("_ ===> _" [51, 51] 50)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     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 {* 
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    15
  {\hfill quote by Larry}\bigskip
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    16
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    17
  \noindent
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    18
  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
    19
  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
    20
  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
    21
  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
    22
  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
    23
  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
    24
  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
    25
  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
    26
  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
    27
  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
    28
  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
    29
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    30
  @{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
    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
  \noindent
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    33
  Similarly one can construct the type of finite sets by quotienting lists
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    34
  according to the equivalence relation
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    35
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    36
  @{text [display] "xs \<approx> ys \<equiv> (\<forall>x. x \<in> xs \<longleftrightarrow> x \<in> ys)"}
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    37
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    38
  \noindent
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    39
  where @{text "\<in>"} stands for membership in a list.
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    40
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    41
  The problem is that in order to start reasoning about, for example integers, 
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    42
  definitions and theorems need to be transferred, or \emph{lifted}, 
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    43
  from the ``raw'' type @{typ "nat \<times> nat"} to the quotient type @{typ int}. 
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    44
  This lifting usually requires a lot of tedious reasoning effort.
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    45
  The purpose of a \emph{quotient package} is to ease the lifting and automate
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    46
  the reasoning involved as much as possible. Such a package is a central
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    47
  component of the new version of Nominal Isabelle where representations 
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    48
  of alpha-equated terms are constructed according to specifications given by
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    49
  the user. 
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    50
  
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    51
  In the context of HOL, there have been several quotient packages (...). The
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    52
  most notable is the one by Homeier (...) implemented in HOL4. However, what is
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    53
  surprising, none of them can deal compositions of quotients, for example with 
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    54
  lifting theorems about @{text "concat"}:
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    55
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    56
  @{text [display] "concat definition"}
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    57
  
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    58
  \noindent
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    59
  One would like to lift this definition to the operation
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    60
  
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    61
  @{text [display] "union definition"}
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    62
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    63
  \noindent
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    64
  What is special about this operation is that we have as input
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    65
  lists of lists which after lifting turn into finite sets of finite
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    66
  sets. 
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    67
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    68
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    69
subsection {* Contributions *}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    70
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    71
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    72
  We present the detailed lifting procedure, which was not shown before.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    73
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    74
  The quotient package presented in this paper has the following
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    75
  advantages over existing packages:
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    76
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    77
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    78
  \item We define quotient composition, function map composition and
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    79
    relation map composition. This lets lifting polymorphic types with
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    80
    subtypes quotiented as well. We extend the notions of
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    81
    respectfullness and preservation to cope with quotient
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    82
    composition.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    83
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    84
  \item We allow lifting only some occurrences of quotiented
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    85
    types. Rsp/Prs extended. (used in nominal)
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    86
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    87
  \item We regularize more thanks to new lemmas. (inductions in
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    88
    nominal).
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    89
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    90
  \item The quotient package is very modular. Definitions can be added
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    91
    separately, rsp and prs can be proved separately and theorems can
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    92
    be lifted on a need basis. (useful with type-classes).
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    93
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    94
  \item Can be used both manually (attribute, separate tactics,
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    95
    rsp/prs databases) and programatically (automated definition of
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    96
    lifted constants, the rsp proof obligations and theorem statement
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    97
    translation according to given quotients).
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    98
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    99
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   100
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   101
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   102
section {* Quotient Type*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   103
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   104
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   105
  Defintion of quotient,
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   106
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   107
  Equivalence,
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   108
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   109
  Relation map and function map
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   110
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   111
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   112
section {* Constants *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   113
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   114
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   115
  Rep and Abs, Rsp and Prs
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   116
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   117
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   118
section {* Lifting Theorems *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   119
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   120
text {* TBD *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   121
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   122
text {* Why providing a statement to prove is necessary is some cases *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   123
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   124
subsection {* Regularization *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   125
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   126
text {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   127
Transformation of the theorem statement:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   128
\begin{itemize}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   129
\item Quantifiers and abstractions involving raw types replaced by bounded ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   130
\item Equalities involving raw types replaced by bounded ones.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   131
\end{itemize}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   132
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   133
The procedure.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   134
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   135
Example of non-regularizable theorem ($0 = 1$).
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   136
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   137
New regularization lemmas:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   138
\begin{lemma}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   139
If @{term R2} is an equivalence relation, then:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   140
\begin{eqnarray}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   141
@{thm (rhs) ball_reg_eqv_range[no_vars]} & = & @{thm (lhs) ball_reg_eqv_range[no_vars]}\\
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   142
@{thm (rhs) bex_reg_eqv_range[no_vars]} & = & @{thm (lhs) bex_reg_eqv_range[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   143
\end{eqnarray}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   144
\end{lemma}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   145
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   146
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   147
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   148
subsection {* Injection *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   149
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   150
subsection {* Cleaning *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   151
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   152
text {* Preservation of quantifiers, abstractions, relations, quotient-constants
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   153
  (definitions) and user given constant preservation lemmas *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   154
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   155
section {* Examples *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   156
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   157
section {* Related Work *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   158
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   159
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   160
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   161
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   162
  \item Peter Homeier's package (and related work from there), John
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   163
    Harrison's one.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   164
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   165
  \item Manually defined quotients in Isabelle/HOL Library (Larry's
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   166
    quotients, Markus's Quotient\_Type, Dixon's FSet, \ldots)
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   167
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   168
  \item Oscar Slotosch defines quotient-type automatically but no
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   169
    lifting.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   170
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   171
  \item PER. And how to avoid it.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   172
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   173
  \item Necessity of Hilbert Choice op.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   174
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   175
  \item Setoids in Coq
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   176
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   177
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   178
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   179
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   180
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   181
end
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   182
(*>*)