Quotient-Paper/Paper.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Thu, 27 May 2010 16:53:12 +0200
changeset 2199 6ce64fb5cbd9
parent 2198 8fe1a706ade7
child 2205 69b4eb4b12c6
permissions -rw-r--r--
qpaper / lemmas used in proofs
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
     1
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     2
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
theory Paper
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
     4
imports "Quotient"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     5
        "LaTeXsugar"
2186
762a739c9eb4 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2183
diff changeset
     6
        "../Nominal/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
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     8
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     9
notation (latex output)
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    10
  rel_conj ("_ OOO _" [53, 53] 52)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    11
and
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    12
  fun_map ("_ ---> _" [51, 51] 50)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    13
and
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    14
  fun_rel ("_ ===> _" [51, 51] 50)
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
    15
and
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
    16
  list_eq (infix "\<approx>" 50) (* Not sure if we want this notation...? *)
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    17
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    18
ML {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    19
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    20
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    21
  let
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    22
    val concl =
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    23
      Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    24
  in
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    25
    case concl of (_ $ l $ r) => proj (l, r)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    26
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    27
  end);
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    28
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    29
setup {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    30
  Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    31
  Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    32
  Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    33
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    34
(*>*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    35
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    36
section {* Introduction *}
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    37
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    38
text {* 
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    39
  {\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
    40
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    41
  \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
    42
  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
    43
  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
    44
  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
    45
  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
    46
  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
    47
  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
    48
  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
    49
  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
    50
  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
    51
  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
    52
  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
    53
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    54
% I would avoid substraction for natural numbers.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    55
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    56
  @{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
    57
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    58
  \noindent
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    59
  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
    60
  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
    61
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    62
  @{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
    63
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    64
  \noindent
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    65
  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
    66
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    67
  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
    68
  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
    69
  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
    70
  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
    71
  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
    72
  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
    73
  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
    74
  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
    75
  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
    76
  
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    77
  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
    78
  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
    79
  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
    80
  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
    81
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
    82
  @{thm [display] concat.simps(1)}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
    83
  @{thm [display] concat.simps(2)[no_vars]}
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
    84
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    85
  \noindent
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
    86
  One would like to lift this definition to the operation:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
    87
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
    88
  @{thm [display] fconcat_empty[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
    89
  @{thm [display] fconcat_insert[no_vars]}
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    90
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    91
  \noindent
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    92
  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
    93
  lists of lists which after lifting turn into finite sets of finite
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
    94
  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
    95
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    96
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    97
subsection {* Contributions *}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    98
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
    99
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   100
  We present the detailed lifting procedure, which was not shown before.
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
  The quotient package presented in this paper has the following
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   103
  advantages over existing packages:
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   104
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   105
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   106
  \item We define quotient composition, function map composition and
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   107
    relation map composition. This lets lifting polymorphic types with
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   108
    subtypes quotiented as well. We extend the notions of
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   109
    respectfulness and preservation to cope with quotient
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   110
    composition.
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
  \item We allow lifting only some occurrences of quotiented
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   113
    types. Rsp/Prs extended. (used in nominal)
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   114
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   115
  \item The quotient package is very modular. Definitions can be added
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   116
    separately, rsp and prs can be proved separately and theorems can
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   117
    be lifted on a need basis. (useful with type-classes). 
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   118
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   119
  \item Can be used both manually (attribute, separate tactics,
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   120
    rsp/prs databases) and programatically (automated definition of
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   121
    lifted constants, the rsp proof obligations and theorem statement
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   122
    translation according to given quotients).
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   123
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   124
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   125
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   126
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   127
section {* Quotient Type*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   128
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   129
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   130
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   131
text {*
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   132
  In this section we present the definitions of a quotient that follow
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   133
  those by Homeier, the proofs can be found there.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   134
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   135
  \begin{definition}[Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   136
  A relation $R$ with an abstraction function $Abs$
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   137
  and a representation function $Rep$ is a \emph{quotient}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   138
  if and only if:
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   139
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   140
  \begin{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   141
  \item @{thm (rhs1) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   142
  \item @{thm (rhs2) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   143
  \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   144
  \end{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   145
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   146
  \end{definition}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   147
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   148
  \begin{definition}[Relation map and function map]\\
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   149
  @{thm fun_rel_def[of "R1" "R2", no_vars]}\\
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   150
  @{thm fun_map_def[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   151
  \end{definition}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   152
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   153
  The main theorems for building higher order quotients is:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   154
  \begin{lemma}[Function Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   155
  If @{thm (prem 1) fun_quotient[no_vars]} and @{thm (prem 2) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   156
  then @{thm (concl) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   157
  \end{lemma}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   158
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   159
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   160
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   161
subsection {* Higher Order Logic *}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   162
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   163
text {*
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   164
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   165
  Types:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   166
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   167
  @{text "\<sigma> ::="} & @{text "\<alpha>"} & \textrm{(type variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   168
      @{text "|"} & @{text "(\<sigma>,\<dots>,\<sigma>)\<kappa>"} & \textrm{(type construction)}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   169
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   170
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   171
  Terms:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   172
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   173
  @{text "t ::="} & @{text "x\<^isup>\<sigma>"} & \textrm{(variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   174
      @{text "|"} & @{text "c\<^isup>\<sigma>"} & \textrm{(constant)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   175
      @{text "|"} & @{text "t t"} & \textrm{(application)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   176
      @{text "|"} & @{text "\<lambda>x\<^isup>\<sigma>. t"} & \textrm{(abstraction)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   177
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   178
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   179
*}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   180
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   181
section {* Constants *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   182
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   183
(* Say more about containers? *)
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   184
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   185
text {*
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   186
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   187
  To define a constant on the lifted type, an aggregate abstraction
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   188
  function is applied to the raw constant. Below we describe the operation
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   189
  that generates
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   190
  an aggregate @{term "Abs"} or @{term "Rep"} function given the
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   191
  compound raw type and the compound quotient type.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   192
  This operation will also be used in translations of theorem statements
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   193
  and in the lifting procedure.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   194
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   195
  The operation is additionally able to descend into types for which
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   196
  maps are known. Such maps for most common types (list, pair, sum,
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   197
  option, \ldots) are described in Homeier, and we assume that @{text "map"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   198
  is the function that returns a map for a given type. Then REP/ABS is defined
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   199
  as follows:
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   200
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   201
  \begin{itemize}
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   202
  \item @{text "ABS(\<alpha>\<^isub>1, \<alpha>\<^isub>2)"} = @{text "id"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   203
  \item @{text "REP(\<alpha>\<^isub>1, \<alpha>\<^isub>2)"}  =  @{text "id"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   204
  \item @{text "ABS(\<sigma>, \<sigma>)"}  =  @{text "id"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   205
  \item @{text "REP(\<sigma>, \<sigma>)"}  =  @{text "id"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   206
  \item @{text "ABS(\<sigma>\<^isub>1\<rightarrow>\<sigma>\<^isub>2,\<tau>\<^isub>1\<rightarrow>\<tau>\<^isub>2)"}  =  @{text "REP(\<sigma>\<^isub>1,\<tau>\<^isub>1) ---> ABS(\<sigma>\<^isub>2,\<tau>\<^isub>2)"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   207
  \item @{text "REP(\<sigma>\<^isub>1\<rightarrow>\<sigma>\<^isub>2,\<tau>\<^isub>1\<rightarrow>\<tau>\<^isub>2)"}  =  @{text "ABS(\<sigma>\<^isub>1,\<tau>\<^isub>1) ---> REP(\<sigma>\<^isub>2,\<tau>\<^isub>2)"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   208
  \item @{text "ABS((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>n))\<kappa>)"}  =  @{text "(map \<kappa>) (ABS(\<sigma>\<^isub>1,\<tau>\<^isub>1)) \<dots> (ABS(\<sigma>\<^isub>n,\<tau>\<^isub>n))"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   209
  \item @{text "REP((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>n))\<kappa>)"}  =  @{text "(map \<kappa>) (REP(\<sigma>\<^isub>1,\<tau>\<^isub>1)) \<dots> (REP(\<sigma>\<^isub>n,\<tau>\<^isub>n))"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   210
  \item @{text "ABS((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>\<^isub>1, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>m))\<kappa>\<^isub>2)"}  =  @{text "Abs_\<kappa>\<^isub>2 \<circ> (map \<kappa>\<^isub>1) (ABS(\<rho>\<^isub>1,\<nu>\<^isub>1) \<dots> (ABS(\<rho>\<^isub>p,\<nu>\<^isub>p)"} provided @{text "\<eta> \<kappa>\<^isub>2 = (\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>p)\<kappa>\<^isub>1 \<and> \<exists>s. s(\<sigma>s\<kappa>\<^isub>1)=\<rho>s\<kappa>\<^isub>1 \<and> s(\<tau>s\<kappa>\<^isub>2)=\<nu>s\<kappa>\<^isub>2"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   211
  \item @{text "REP((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>\<^isub>1, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>m))\<kappa>\<^isub>2)"}  =  @{text "(map \<kappa>\<^isub>1) (REP(\<rho>\<^isub>1,\<nu>\<^isub>1) \<dots> (REP(\<rho>\<^isub>p,\<nu>\<^isub>p) \<circ> Rep_\<kappa>\<^isub>2"} provided @{text "\<eta> \<kappa>\<^isub>2 = (\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>p)\<kappa>\<^isub>1 \<and> \<exists>s. s(\<sigma>s\<kappa>\<^isub>1)=\<rho>s\<kappa>\<^isub>1 \<and> s(\<tau>s\<kappa>\<^isub>2)=\<nu>s\<kappa>\<^isub>2"}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   212
  \end{itemize}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   213
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   214
  Apart from the last 2 points the definition is same as the one implemented in
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   215
  in Homeier's HOL package, below is the definition of @{term fconcat}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   216
  that shows the last points:
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   217
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   218
  @{thm fconcat_def[no_vars]}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   219
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   220
  The aggregate @{term Abs} function takes a finite set of finite sets
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   221
  and applies @{term "map rep_fset"} composed with @{term rep_fset} to
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   222
  its input, obtaining a list of lists, passes the result to @{term concat}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   223
  obtaining a list and applies @{term abs_fset} obtaining the composed
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   224
  finite set.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   225
*}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   226
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   227
subsection {* Respectfulness *}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   228
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   229
text {*
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   230
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   231
  A respectfulness lemma for a constant states that the equivalence
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   232
  class returned by this constant depends only on the equivalence
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   233
  classes of the arguments applied to the constant. This can be
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   234
  expressed in terms of an aggregate relation between the constant
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   235
  and itself, for example the respectfullness for @{term "append"}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   236
  can be stated as:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   237
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   238
  @{thm [display] append_rsp[no_vars]}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   239
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   240
  \noindent
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   241
  Which is equivalent to:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   242
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   243
  @{thm [display] append_rsp_unfolded[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   244
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   245
  Below we show the algorithm for finding the aggregate relation.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   246
  This algorithm uses
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   247
  the relation composition which we define as:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   248
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   249
  \begin{definition}[Composition of Relations]
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   250
  @{abbrev "rel_conj R1 R2"} where @{text OO} is the predicate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   251
  composition @{thm pred_compI[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   252
  \end{definition}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   253
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   254
  Given an aggregate raw type and quotient type:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   255
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   256
  \begin{itemize}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   257
  \item For equal types or free type variables return equality
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   258
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   259
  \item For equal type constructors use the appropriate rel
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   260
    function applied to the results for the argument pairs
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   261
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   262
  \item For unequal type constructors, look in the quotients information
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   263
    for a quotient type that matches the type constructor, and instantiate
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   264
    the type appropriately getting back an instantiation environment. We
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   265
    apply the environment to the arguments and recurse composing it with
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   266
    the aggregate relation function.
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   267
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   268
  \end{itemize}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   269
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   270
  Again, the the behaviour of our algorithm in the last situation is
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   271
  novel, so lets look at the example of respectfullness for @{term concat}.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   272
  The statement as computed by the algorithm above is:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   273
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   274
  @{thm [display] concat_rsp[no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   275
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   276
  \noindent
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   277
  By unfolding the definition of relation composition and relation map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   278
  we can see the equivalent statement just using the primitive list
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   279
  equivalence relation:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   280
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   281
  @{thm [display] concat_rsp_unfolded[of "a" "a'" "b'" "b", no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   282
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   283
  The statement reads that, for any lists of lists @{term a} and @{term b}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   284
  if there exist intermediate lists of lists @{term "a'"} and @{term "b'"}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   285
  such that each element of @{term a} is in the relation with an appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   286
  element of @{term a'}, @{term a'} is in relation with @{term b'} and each
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   287
  element of @{term b'} is in relation with the appropriate element of
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   288
  @{term b}.
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   289
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   290
*}
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   291
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   292
subsection {* Preservation *}
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   293
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   294
text {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   295
  To be able to lift theorems that talk about constants that are not
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   296
  lifted but whose type changes when lifting is performed additionally
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   297
  preservation theorems are needed.
2196
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   298
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   299
  To lift theorems that talk about insertion in lists of lifted types
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   300
  we need to know that for any quotient type with the abstraction and
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   301
  representation functions @{text "Abs"} and @{text Rep} we have:
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   302
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   303
  @{thm [display] (concl) cons_prs[no_vars]}
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   304
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   305
  This is not enough to lift theorems that talk about quotient compositions.
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   306
  For some constants (for example empty list) it is possible to show a
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   307
  general compositional theorem, but for @{term "op #"} it is necessary
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   308
  to show that it respects the particular quotient type:
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   309
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   310
  @{thm [display] insert_preserve2[no_vars]}
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   311
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   312
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   313
subsection {* Composition of Quotient theorems *}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   314
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   315
text {*
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   316
  Given two quotients, one of which quotients a container, and the
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   317
  other quotients the type in the container, we can write the
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   318
  composition of those quotients. To compose two quotient theorems
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   319
  we compose the relations with relation composition
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   320
  and the abstraction and relation functions with function composition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   321
  The @{term "Rep"} and @{term "Abs"} functions that we obtain are
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   322
  the same as the ones created by in the aggregate functions and the
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   323
  relation is the same as the one given by aggregate relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   324
  This becomes especially interesting
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   325
  when we compose the quotient with itself, as there is no simple
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   326
  intermediate step.
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   327
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   328
  Lets take again the example of @{term concat}. To be able to lift
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   329
  theorems that talk about it we will first prove the composition
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   330
  quotient theorems, which then lets us perform the lifting procedure
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   331
  in an unchanged way:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   332
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   333
  @{thm [display] quotient_compose_list[no_vars]}
2192
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   334
*}
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   335
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   336
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   337
section {* Lifting Theorems *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   338
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   339
text {*
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   340
  The core of the quotient package takes an original theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   341
  talks about the raw types, and the statement of the theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   342
  it is supposed to produce. This is different from other existing
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   343
  quotient packages, where only the raw theorems was necessary.
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   344
  We notice that in some cases only some occurrences of the raw
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   345
  types need to be lifted. This is for example the case in the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   346
  new Nominal package, where a raw datatype that talks about
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   347
  pairs of natural numbers or strings (being lists of characters)
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   348
  should not be changed to a quotient datatype with constructors
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   349
  taking integers or finite sets of characters. To simplify the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   350
  use of the quotient package we additionally provide an automated
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   351
  statement translation mechanism that replaces occurrences of
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   352
  types that match given quotients by appropriate lifted types.
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   353
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   354
  Lifting the theorems is performed in three steps. In the following
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   355
  we call these steps \emph{regularization}, \emph{injection} and
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   356
  \emph{cleaning} following the names used in Homeier's HOL
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   357
  implementation.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   358
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   359
  We first define the statement of the regularized theorem based
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   360
  on the original theorem and the goal theorem. Then we define
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   361
  the statement of the injected theorem, based on the regularized
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   362
  theorem and the goal. We then show the 3 proofs, and all three
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   363
  can be performed independently from each other.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   364
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   365
*}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   366
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   367
subsection {* Regularization and Injection statements *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   368
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   369
text {*
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   370
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   371
  The function that gives the statement of the regularized theorem
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   372
  takes the statement of the raw theorem (a term) and the statement
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   373
  of the lifted theorem. The intuition behind the procedure is that
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   374
  it replaces quantifiers and abstractions involving raw types
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   375
  by bounded ones, and equalities involving raw types are replaced
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   376
  by appropriate aggregate relations. It is defined as follows:
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   377
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   378
  \begin{itemize}
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   379
  \item @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<sigma>. s) = \<lambda>x : \<sigma>. REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   380
  \item @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<tau>. s) = \<lambda>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   381
  \item @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<sigma>. s) = \<forall>x : \<sigma>. REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   382
  \item @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<tau>. s) = \<forall>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   383
  \item @{text "REG ((op =) : \<sigma>, (op =) : \<sigma>) = (op =) : \<sigma>"}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   384
  \item @{text "REG ((op =) : \<sigma>, (op =) : \<tau>) = REL (\<sigma>, \<tau>) : \<sigma>"}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   385
  \item @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   386
  \item @{text "REG (v\<^isub>1, v\<^isub>2) = v\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   387
  \item @{text "REG (c\<^isub>1, c\<^isub>2) = c\<^isub>1"}
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   388
  \end{itemize}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   389
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   390
  Existential quantifiers and unique existential quantifiers are defined
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   391
  similarily to the universal one.
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   392
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   393
  The function that gives the statment of the injected theorem
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   394
  takes the statement of the regularized theorems and the statement
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   395
  of the lifted theorem both as terms.
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   396
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   397
  \begin{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   398
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<sigma>) = \<lambda>x. (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   399
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   400
  \item @{text "INJ ((\<lambda>x \<in> R. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x \<in> R. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   401
  \item @{text "INJ (\<forall> t, \<forall> s) = \<forall> (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   402
  \item @{text "INJ (\<forall> t \<in> R, \<forall> s) = \<forall> (INJ (t, s) \<in> R"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   403
  \item @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   404
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<sigma>) = v\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   405
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (v\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   406
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<sigma>) = c\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   407
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (c\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   408
  \end{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   409
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   410
  For existential quantifiers and unique existential quantifiers it is
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   411
  defined similarily to the universal one.
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   412
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   413
*}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   414
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   415
subsection {* Proof of Regularization *}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   416
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   417
text {*
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   418
  Example of non-regularizable theorem ($0 = 1$).
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   419
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   420
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   421
  Separtion of regularization from injection thanks to the following 2 lemmas:
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   422
  \begin{lemma}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   423
  If @{term R2} is an equivalence relation, then:
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   424
  \begin{eqnarray}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   425
  @{thm (rhs) ball_reg_eqv_range[no_vars]} & = & @{thm (lhs) ball_reg_eqv_range[no_vars]}\\
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   426
  @{thm (rhs) bex_reg_eqv_range[no_vars]} & = & @{thm (lhs) bex_reg_eqv_range[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   427
  \end{eqnarray}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   428
  \end{lemma}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   429
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   430
  Other lemmas used in regularization:
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   431
  @{thm [display] ball_reg_eqv[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   432
  @{thm [display] bex_reg_eqv[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   433
  @{thm [display] babs_reg_eqv[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   434
  @{thm [display] babs_simp[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   435
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   436
  @{thm [display] ball_reg_right[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   437
  @{thm [display] bex_reg_left[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   438
  @{thm [display] bex1_bexeq_reg[no_vars]}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   439
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   440
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   441
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   442
subsection {* Injection *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   443
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   444
text {*
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   445
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   446
  The 2 key lemmas are:
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   447
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   448
  @{thm [display] apply_rsp[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   449
  @{thm [display] rep_abs_rsp[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   450
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   451
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   452
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   453
*}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   454
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   455
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   456
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   457
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   458
subsection {* Cleaning *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   459
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   460
text {* Preservation of quantifiers, abstractions, relations, quotient-constants
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   461
  (definitions) and user given constant preservation lemmas *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   462
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   463
section {* Examples *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   464
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   465
section {* Related Work *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   466
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   467
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   468
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   469
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   470
  \item Peter Homeier's package~\cite{Homeier05} (and related work from there)
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   471
  \item John Harrison's one~\cite{harrison-thesis} is the first one to lift theorems
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   472
    but only first order.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   473
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   474
  \item PVS~\cite{PVS:Interpretations}
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   475
  \item MetaPRL~\cite{Nogin02}
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   476
  \item Manually defined quotients in Isabelle/HOL Library (Markus's Quotient\_Type,
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   477
    Dixon's FSet, \ldots)
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   478
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   479
  \item Oscar Slotosch defines quotient-type automatically but no
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   480
    lifting~\cite{Slotosch97}.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   481
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   482
  \item PER. And how to avoid it.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   483
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   484
  \item Necessity of Hilbert Choice op and Larry's quotients~\cite{Paulson06}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   485
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   486
  \item Setoids in Coq and \cite{ChicliPS02}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   487
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   488
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   489
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   490
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   491
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   492
end
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   493
(*>*)