Quotient-Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Wed, 16 Jun 2010 22:29:42 +0100
changeset 2278 337569f85398
parent 2277 816204c76e90
child 2279 2cbcdaba795a
permissions -rw-r--r--
conclusion done
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
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
     3
imports "Quotient"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
        "LaTeXsugar"
2186
762a739c9eb4 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2183
diff changeset
     5
        "../Nominal/FSet"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
begin
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     7
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     8
notation (latex output)
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
     9
  rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
    10
  pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
    11
  "op -->" (infix "\<longrightarrow>" 100) and
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
    12
  "==>" (infix "\<Longrightarrow>" 100) and
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    13
  fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    14
  fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    15
  list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    16
  fempty ("\<emptyset>") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    17
  funion ("_ \<union> _") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    18
  finsert ("{_} \<union> _") and 
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    19
  Cons ("_::_") and
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    20
  concat ("flat") and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    21
  fconcat ("\<Union>")
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    22
 
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    23
  
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    24
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    25
ML {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    26
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
    27
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    28
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    29
  let
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    30
    val concl =
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    31
      Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    32
  in
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    33
    case concl of (_ $ l $ r) => proj (l, r)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    34
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    35
  end);
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    36
*}
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
    37
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    38
setup {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    39
  Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    40
  Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    41
  Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    42
*}
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
    43
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    44
(*>*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    46
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    47
section {* Introduction *}
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    48
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    49
text {* 
2205
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    50
   \begin{flushright}
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    51
  {\em ``Not using a [quotient] package has its advantages: we do not have to\\ 
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    52
    collect all the theorems we shall ever want into one giant list;''}\\
2213
231a20534950 improved abstract, some tuning
Christian Urban <urbanc@in.tum.de>
parents: 2212
diff changeset
    53
    Larry Paulson \cite{Paulson06}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    54
  \end{flushright}
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    55
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    56
  \noindent
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    57
  Isabelle is a popular generic theorem prover in which many logics can be
2214
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    58
  implemented. The most widely used one, however, is Higher-Order Logic
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    59
  (HOL). This logic consists of a small number of axioms and inference rules
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    60
  over a simply-typed term-language. Safe reasoning in HOL is ensured by two
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    61
  very restricted mechanisms for extending the logic: one is the definition of
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    62
  new constants in terms of existing ones; the other is the introduction of
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    63
  new types by identifying non-empty subsets in existing types. It is well
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    64
  understood how to use both mechanisms for dealing with quotient
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    65
  constructions in HOL (see \cite{Homeier05,Paulson06}).  For example the
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    66
  integers in Isabelle/HOL are constructed by a quotient construction over the
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    67
  type @{typ "nat \<times> nat"} and the equivalence relation
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    68
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    69
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    70
  @{text "(n\<^isub>1, n\<^isub>2) \<approx> (m\<^isub>1, m\<^isub>2) \<equiv> n\<^isub>1 + m\<^isub>2 = m\<^isub>1 + n\<^isub>2"}\hfill\numbered{natpairequiv}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    71
  \end{isabelle}
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    72
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    73
  \noindent
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    74
  This constructions yields the new type @{typ int} and definitions for @{text
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    75
  "0"} and @{text "1"} of type @{typ int} can be given in terms of pairs of
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    76
  natural numbers (namely @{text "(0, 0)"} and @{text "(1, 0)"}). Operations
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    77
  such as @{text "add"} with type @{typ "int \<Rightarrow> int \<Rightarrow> int"} can be defined in
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    78
  terms of operations on pairs of natural numbers (namely @{text
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    79
  "add_pair (n\<^isub>1, m\<^isub>1) (n\<^isub>2,
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
    80
  m\<^isub>2) \<equiv> (n\<^isub>1 + n\<^isub>2, m\<^isub>1 + m\<^isub>2)"}).
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
    81
  Similarly one can construct the type of finite sets, written @{term "\<alpha> fset"}, 
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    82
  by quotienting the type @{text "\<alpha> list"} according to the equivalence relation
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    83
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    84
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    85
  @{text "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}\hfill\numbered{listequiv}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    86
  \end{isabelle}
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    87
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    88
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    89
  which states that two lists are equivalent if every element in one list is
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    90
  also member in the other. The empty finite set, written @{term "{||}"}, can
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    91
  then be defined as the empty list and the union of two finite sets, written
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    92
  @{text "\<union>"}, as list append.
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    93
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
    94
  Quotients are important in a variety of areas, but they are really ubiquitous in
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
    95
  the area of reasoning about programming language calculi. A simple example
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
    96
  is the lambda-calculus, whose raw terms are defined as
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
    97
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    98
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
    99
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   100
  @{text "t ::= x | t t | \<lambda>x.t"}\hfill\numbered{lambda}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   101
  \end{isabelle}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   102
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   103
  \noindent
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   104
  The problem with this definition arises, for instance, when one attempts to
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   105
  prove formally the substitution lemma \cite{Barendregt81} by induction
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   106
  over the structure of terms. This can be fiendishly complicated (see
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   107
  \cite[Pages 94--104]{CurryFeys58} for some ``rough'' sketches of a proof
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   108
  about raw lambda-terms). In contrast, if we reason about
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   109
  $\alpha$-equated lambda-terms, that means terms quotient according to
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   110
  $\alpha$-equivalence, then the reasoning infrastructure provided, 
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   111
  for example, by Nominal Isabelle \cite{UrbanKaliszyk11} makes the formal 
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   112
  proof of the substitution lemma almost trivial. 
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   113
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   114
  The difficulty is that in order to be able to reason about integers, finite
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   115
  sets or $\alpha$-equated lambda-terms one needs to establish a reasoning
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   116
  infrastructure by transferring, or \emph{lifting}, definitions and theorems
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   117
  from the raw type @{typ "nat \<times> nat"} to the quotient type @{typ int}
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   118
  (similarly for finite sets and $\alpha$-equated lambda-terms). This lifting
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   119
  usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.  
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   120
  It is feasible to do this work manually, if one has only a few quotient
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   121
  constructions at hand. But if they have to be done over and over again, as in 
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   122
  Nominal Isabelle, then manual reasoning is not an option.
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   123
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   124
  The purpose of a \emph{quotient package} is to ease the lifting of theorems
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   125
  and automate the reasoning as much as possible. In the
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   126
  context of HOL, there have been a few quotient packages already
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   127
  \cite{harrison-thesis,Slotosch97}. The most notable one is by Homeier
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   128
  \cite{Homeier05} implemented in HOL4.  The fundamental construction these
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   129
  quotient packages perform can be illustrated by the following picture:
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   130
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   131
  \begin{center}
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   132
  \mbox{}\hspace{20mm}\begin{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   133
  %%\draw[step=2mm] (-4,-1) grid (4,1);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   134
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   135
  \draw[very thick] (0.7,0.3) circle (4.85mm);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   136
  \draw[rounded corners=1mm, very thick] ( 0.0,-0.9) rectangle ( 1.8, 0.9);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   137
  \draw[rounded corners=1mm, very thick] (-1.95,0.8) rectangle (-2.9,-0.195);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   138
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   139
  \draw (-2.0, 0.8) --  (0.7,0.8);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   140
  \draw (-2.0,-0.195)  -- (0.7,-0.195);
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   141
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   142
  \draw ( 0.7, 0.23) node {\begin{tabular}{@ {}c@ {}}equiv-\\[-1mm]clas.\end{tabular}};
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   143
  \draw (-2.45, 0.35) node {\begin{tabular}{@ {}c@ {}}new\\[-1mm]type\end{tabular}};
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   144
  \draw (1.8, 0.35) node[right=-0.1mm]
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   145
    {\begin{tabular}{@ {}l@ {}}existing\\[-1mm] type\\ (sets of raw elements)\end{tabular}};
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   146
  \draw (0.9, -0.55) node {\begin{tabular}{@ {}l@ {}}non-empty\\[-1mm]subset\end{tabular}};
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   147
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   148
  \draw[->, very thick] (-1.8, 0.36) -- (-0.1,0.36);
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   149
  \draw[<-, very thick] (-1.8, 0.16) -- (-0.1,0.16);
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   150
  \draw (-0.95, 0.26) node[above=0.4mm] {@{text Rep}};
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   151
  \draw (-0.95, 0.26) node[below=0.4mm] {@{text Abs}};
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   152
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   153
  \end{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   154
  \end{center}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   155
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   156
  \noindent
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   157
  The starting point is an existing type, to which we refer as the
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   158
  \emph{raw type} and over which an equivalence relation given by the user is
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   159
  defined. With this input the package introduces a new type, to which we
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   160
  refer as the \emph{quotient type}. This type comes with an
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   161
  \emph{abstraction} and a \emph{representation} function, written @{text Abs}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   162
  and @{text Rep}.\footnote{Actually slightly more basic functions are given;
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   163
  the functions @{text Abs} and @{text Rep} need to be derived from them. We
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   164
  will show the details later. } They relate elements in the
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   165
  existing type to elements in the new type and vice versa, and can be uniquely
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   166
  identified by their quotient type. For example for the integer quotient construction
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   167
  the types of @{text Abs} and @{text Rep} are
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   168
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   169
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   170
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   171
  @{text "Abs :: nat \<times> nat \<Rightarrow> int"}\hspace{10mm}@{text "Rep :: int \<Rightarrow> nat \<times> nat"}
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   172
  \end{isabelle}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   173
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   174
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   175
  We therefore often write @{text Abs_int} and @{text Rep_int} if the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   176
  typing information is important. 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   177
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   178
  Every abstraction and representation function stands for an isomorphism
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   179
  between the non-empty subset and elements in the new type. They are
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   180
  necessary for making definitions involving the new type. For example @{text
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   181
  "0"} and @{text "1"} of type @{typ int} can be defined as
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   182
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   183
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   184
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   185
  @{text "0 \<equiv> Abs_int (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs_int (1, 0)"}
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   186
  \end{isabelle}
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   187
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   188
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   189
  Slightly more complicated is the definition of @{text "add"} having type 
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   190
  @{typ "int \<Rightarrow> int \<Rightarrow> int"}. Its definition is as follows
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   191
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   192
   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   193
  @{text "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   194
  \hfill\numbered{adddef}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   195
  \end{isabelle}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   196
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   197
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   198
  where we take the representation of the arguments @{text n} and @{text m},
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   199
  add them according to the function @{text "add_pair"} and then take the
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   200
  abstraction of the result.  This is all straightforward and the existing
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   201
  quotient packages can deal with such definitions. But what is surprising
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   202
  that none of them can deal with slightly more complicated definitions involving
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   203
  \emph{compositions} of quotients. Such compositions are needed for example
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   204
  in case of quotienting lists to yield finite sets and the operator that 
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   205
  flattens lists of lists, defined as follows
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
   206
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   207
  @{thm [display, indent=10] concat.simps(1) concat.simps(2)[no_vars]}
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
   208
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   209
  \noindent
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   210
  We expect that the corresponding operator on finite sets, written @{term "fconcat"},
2248
Christian Urban <urbanc@in.tum.de>
parents: 2247 2246
diff changeset
   211
  builds finite unions of finite sets:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   212
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   213
  @{thm [display, indent=10] fconcat_empty[no_vars] fconcat_insert[no_vars]}
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
   214
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   215
  \noindent
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   216
  The quotient package should automatically provide us with a definition for @{text "\<Union>"} in
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   217
  terms of @{text flat}, @{text Rep_fset} and @{text Abs_fset}. The problem is 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   218
  that the method  used in the existing quotient
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   219
  packages of just taking the representation of the arguments and then taking
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   220
  the abstraction of the result is \emph{not} enough. The reason is that in case
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   221
  of @{text "\<Union>"} we obtain the incorrect definition
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   222
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   223
  @{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat (Rep_fset S))"}
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   224
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   225
  \noindent
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   226
  where the right-hand side is not even typable! This problem can be remedied in the
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   227
  existing quotient packages by introducing an intermediate step and reasoning
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   228
  about flattening of lists of finite sets. However, this remedy is rather
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   229
  cumbersome and inelegant in light of our work, which can deal with such
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   230
  definitions directly. The solution is that we need to build aggregate
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   231
  representation and abstraction functions, which in case of @{text "\<Union>"}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   232
  generate the following definition
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   233
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   234
  @{text [display, indent=10] "\<Union> S \<equiv> Abs_fset (flat ((map Rep_fset \<circ> Rep_fset) S))"}
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   235
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   236
  \noindent
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   237
  where @{term map} is the usual mapping function for lists. In this paper we
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   238
  will present a formal definition of our aggregate abstraction and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   239
  representation functions (this definition was omitted in \cite{Homeier05}). 
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   240
  They generate definitions, like the one above for @{text "\<Union>"}, 
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   241
  according to the type of the raw constant and the type
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   242
  of the quotient constant. This means we also have to extend the notions
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   243
  of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   244
  from Homeier \cite{Homeier05}.
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   245
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   246
  In addition we are able to address the criticism by Paulson \cite{Paulson06} cited
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   247
  at the beginning of this section about having to collect theorems that are
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   248
  lifted from the raw level to the quotient level into one giant list. Our
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   249
  quotient package is the first one that is modular so that it allows to lift
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   250
  single theorems separately. This has the advantage for the user of being able to develop a
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   251
  formal theory interactively as a natural progression. A pleasing side-result of
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   252
  the modularity is that we are able to clearly specify what is involved
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   253
  in the lifting process (this was only hinted at in \cite{Homeier05} and
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   254
  implemented as a ``rough recipe'' in ML-code).
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   255
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   256
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   257
  The paper is organised as follows: Section \ref{sec:prelims} presents briefly
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   258
  some necessary preliminaries; Section \ref{sec:type} describes the definitions 
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   259
  of quotient types and shows how definitions of constants can be made over 
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   260
  quotient types. Section \ref{sec:resp} introduces the notions of respectfulness
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   261
  and preservation; Section \ref{sec:lift} describes the lifting of theorems;
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   262
  Section \ref{sec:examples} presents some examples
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   263
  and Section \ref{sec:conc} concludes and compares our results to existing 
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   264
  work.
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
   265
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   266
2257
Christian Urban <urbanc@in.tum.de>
parents: 2256
diff changeset
   267
section {* Preliminaries and General Quotients\label{sec:prelims} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   268
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   269
text {*
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   270
  We give in this section a crude overview of HOL and describe the main
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   271
  definitions given by Homeier for quotients \cite{Homeier05}.
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   272
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   273
  At its core, HOL is based on a simply-typed term language, where types are 
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   274
  recorded in Church-style fashion (that means, we can always infer the type of 
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   275
  a term and its subterms without any additional information). The grammars
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   276
  for types and terms are as follows
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   277
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   278
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   279
  \begin{tabular}{@ {}rl@ {\hspace{3mm}}l@ {}}
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   280
  @{text "\<sigma>, \<tau> ::="} & @{text "\<alpha> | (\<sigma>,\<dots>, \<sigma>) \<kappa>"} & (type variables and type constructors)\\
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   281
  @{text "t, s ::="} & @{text "x\<^isup>\<sigma> | c\<^isup>\<sigma> | t t | \<lambda>x\<^isup>\<sigma>. t"} & 
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   282
  (variables, constants, applications and abstractions)\\
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   283
  \end{tabular}
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   284
  \end{isabelle}
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   285
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   286
  \noindent
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   287
  We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   288
  @{text "\<sigma>s"} to stand for collections of type variables and types,
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   289
  respectively.  The type of a term is often made explicit by writing @{text
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   290
  "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   291
  type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   292
  constants; a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   293
  bool"}, and the identity function with type @{text "id :: \<sigma> => \<sigma>"} is
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   294
  defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}).
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   295
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   296
  An important point to note is that theorems in HOL can be seen as a subset
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   297
  of terms that are constructed specially (namely through axioms and prove
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   298
  rules). As a result we are able to define automatic proof
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   299
  procedures showing that one theorem implies another by decomposing the term
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   300
  underlying the first theorem.
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   301
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   302
  Like Homeier, our work relies on map-functions defined for every type
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   303
  constructor taking some arguments, for example @{text map} for lists. Homeier
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   304
  describes in \cite{Homeier05} map-functions for products, sums, options and
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   305
  also the following map for function types
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   306
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   307
  @{thm [display, indent=10] fun_map_def[no_vars, THEN eq_reflection]}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   308
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   309
  \noindent
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   310
  Using this map-function, we can give the following, equivalent, but more 
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   311
  uniform, definition for @{text add} shown in \eqref{adddef}:
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   312
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   313
  @{text [display, indent=10] "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   314
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   315
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   316
  Using extensionality and unfolding the definition of @{text "\<singlearr>"}, 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   317
  we can get back to \eqref{adddef}. 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   318
  In what follows we shall use the convention to write @{text "map_\<kappa>"} for a map-function 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   319
  of the type-constructor @{text \<kappa>}. In our implementation we maintain 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   320
  a database of these map-functions that can be dynamically extended.
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   321
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   322
  It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   323
  which define equivalence relations in terms of constituent equivalence
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   324
  relations. For example given two equivalence relations @{text "R\<^isub>1"}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   325
  and @{text "R\<^isub>2"}, we can define an equivalence relations over 
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   326
  products as follows
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   327
  %
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   328
  @{text [display, indent=10] "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   329
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   330
  \noindent
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   331
  Homeier gives also the following operator for defining equivalence 
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   332
  relations over function types
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   333
  %
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   334
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   335
  @{thm fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   336
  \hfill\numbered{relfun}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   337
  \end{isabelle}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   338
  
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   339
  \noindent
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   340
  In the context of quotients, the following two notions from are \cite{Homeier05} 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   341
  needed later on.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   342
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   343
  \begin{definition}[Respects]\label{def:respects}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   344
  An element @{text "x"} respects a relation @{text "R"} provided @{text "R x x"}.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   345
  \end{definition}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   346
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   347
  \begin{definition}[Bounded Quantification and Bounded Abstractions]\label{def:babs}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   348
  @{text "\<forall>x \<in> S. P x"} holds if for all @{text x}, @{text "x \<in> S"} implies @{text "P x"};
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   349
  and @{text "(\<lambda>x \<in> S. f x) = f x"} provided @{text "x \<in> S"}.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   350
  \end{definition}
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   351
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   352
  The central definition in Homeier's work \cite{Homeier05} relates equivalence 
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   353
  relations, abstraction and representation functions:
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   354
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   355
  \begin{definition}[Quotient Types]
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   356
  Given a relation $R$, an abstraction function $Abs$
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   357
  and a representation function $Rep$, the predicate @{term "Quotient R Abs Rep"}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   358
  means
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   359
  \begin{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   360
  \item @{thm (rhs1) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   361
  \item @{thm (rhs2) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   362
  \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   363
  \end{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   364
  \end{definition}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   365
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   366
  \noindent
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   367
  The value of this definition is that validity of @{text "Quotient R Abs Rep"} can 
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   368
  often be proved in terms of the validity of @{text "Quotient"} over the constituent 
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   369
  types of @{text "R"}, @{text Abs} and @{text Rep}. 
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   370
  For example Homeier proves the following property for higher-order quotient
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   371
  types:
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   372
 
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   373
  \begin{proposition}\label{funquot}
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   374
  @{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2" 
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   375
      and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   376
  \end{proposition}
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   377
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   378
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   379
  As a result, Homeier is able to build an automatic prover that can nearly
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   380
  always discharge a proof obligation involving @{text "Quotient"}. Our quotient
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   381
  package makes heavy 
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   382
  use of this part of Homeier's work including an extension 
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   383
  to deal with compositions of equivalence relations defined as follows:
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   384
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   385
  \begin{definition}[Composition of Relations]
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   386
  @{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   387
  composition defined by @{thm (concl) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   388
  holds if and only if @{thm (prem 1) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]} and
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   389
  @{thm (prem 2) pred_compI[of "R\<^isub>1" "x" "y" "R\<^isub>2" "z"]}.
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   390
  \end{definition}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   391
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   392
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   393
  Unfortunately, there are two predicaments with compositions of relations.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   394
  First, a general quotient theorem, like the one given in Proposition \ref{funquot},
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   395
  cannot be stated inside HOL, because of the restriction on types.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   396
  Second, even if we were able to state such a quotient theorem, it
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   397
  would not be true in general. However, we can prove specific and useful 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   398
  instances of the quotient theorem. We will 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   399
  show an example in Section \ref{sec:resp}.
2268
1fd6809f5a44 Definition of Respects.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2267
diff changeset
   400
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   401
*}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   402
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   403
section {* Quotient Types and Quotient Definitions\label{sec:type} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   404
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   405
text {*
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   406
  The first step in a quotient construction is to take a name for the new
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   407
  type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   408
  defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   409
  relation must be @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   410
  the quotient type declaration is therefore
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   411
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   412
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   413
  \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}\hfill\numbered{typedecl}
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   414
  \end{isabelle}
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   415
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   416
  \noindent
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   417
  and a proof that @{text "R"} is indeed an equivalence relation. Two concrete
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   418
  examples are
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   419
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   420
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   421
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   422
  \begin{tabular}{@ {}l}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   423
  \isacommand{quotient\_type}~~@{text "int = nat \<times> nat / \<approx>\<^bsub>nat \<times> nat\<^esub>"}\\
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   424
  \isacommand{quotient\_type}~~@{text "\<alpha> fset = \<alpha> list / \<approx>\<^bsub>list\<^esub>"}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   425
  \end{tabular}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   426
  \end{isabelle}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   427
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   428
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   429
  which introduce the type of integers and of finite sets using the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   430
  equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   431
  "\<approx>\<^bsub>list\<^esub>"} defined in \eqref{natpairequiv} and
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   432
  \eqref{listequiv}, respectively (the proofs about being equivalence
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   433
  relations is omitted).  Given this data, we define for declarations shown in
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   434
  \eqref{typedecl} the quotient types internally as
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   435
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   436
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   437
  \isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   438
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   439
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   440
  \noindent
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   441
  where the right-hand side is the (non-empty) set of equivalence classes of
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   442
  @{text "R"}. The constraint in this declaration is that the type variables
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   443
  in the raw type @{text "\<sigma>"} must be included in the type variables @{text
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   444
  "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will then provide us with the following
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   445
  abstraction and representation functions 
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   446
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   447
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   448
  @{text "abs_\<kappa>\<^isub>q :: \<sigma> set \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}\hspace{10mm}@{text "rep_\<kappa>\<^isub>q :: \<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma> set"}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   449
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   450
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   451
  \noindent 
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   452
  As can be seen from the type, they relate the new quotient type and equivalence classes of the raw
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   453
  type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   454
  to work with the following derived abstraction and representation functions
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   455
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   456
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   457
  @{text "Abs_\<kappa>\<^isub>q x \<equiv> abs_\<kappa>\<^isub>q (R x)"}\hspace{10mm}@{text "Rep_\<kappa>\<^isub>q x \<equiv> \<epsilon> (rep_\<kappa>\<^isub>q x)"}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   458
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   459
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   460
  \noindent
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   461
  on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"} in the
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   462
  definition of @{text "Rep_\<kappa>\<^isub>q"}. These derived notions relate the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   463
  quotient type and the raw type directly, as can be seen from their type,
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   464
  namely @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"},
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   465
  respectively.  Given that @{text "R"} is an equivalence relation, the
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   466
  following property holds  for every quotient type 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   467
  (for the proof see \cite{Homeier05}).
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   468
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   469
  \begin{proposition}
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   470
  @{text "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   471
  \end{proposition}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   472
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   473
  The next step in a quotient construction is to introduce definitions of new constants
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   474
  involving the quotient type. These definitions need to be given in terms of concepts
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   475
  of the raw type (remember this is the only way how to extend HOL
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   476
  with new definitions). For the user the visible part of such definitions is the declaration
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   477
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   478
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   479
  \isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   480
  \end{isabelle}
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   481
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   482
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   483
  where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   484
  and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   485
  given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ 
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   486
  in places where a quotient and raw type is involved). Two concrete examples are
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   487
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   488
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   489
  \begin{tabular}{@ {}l}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   490
  \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   491
  \isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   492
  \isacommand{is}~~@{text "flat"} 
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   493
  \end{tabular}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   494
  \end{isabelle}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   495
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   496
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   497
  The first one declares zero for integers and the second the operator for
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   498
  building unions of finite sets (@{text "flat"} having the type 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   499
  @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"}). 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   500
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   501
  The problem for us is that from such declarations we need to derive proper
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   502
  definitions using the @{text "Abs"} and @{text "Rep"} functions for the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   503
  quotient types involved. The data we rely on is the given quotient type
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   504
  @{text "\<tau>"} and the raw type @{text "\<sigma>"}.  They allow us to define \emph{aggregate
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   505
  abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   506
  \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we give below. The idea behind
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   507
  these two functions is to simultaneously descend into the raw types @{text \<sigma>} and 
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   508
  quotient types @{text \<tau>}, and generate the appropriate
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   509
  @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
2269
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   510
  we generate just the identity whenever the types are equal. On the ``way'' down,
e4699a240d2c tuned everytinh up to section 4
Christian Urban <urbanc@in.tum.de>
parents: 2265
diff changeset
   511
  however we might have to use map-functions to let @{text Abs} and @{text Rep} act
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   512
  over the appropriate types. In what follows we use the short-hand notation 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   513
  @{text "ABS (\<sigma>s, \<tau>s)"} to mean @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1)\<dots>ABS (\<sigma>\<^isub>i, \<tau>\<^isub>i)"}; similarly 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   514
  for @{text REP}.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   515
  %
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   516
  \begin{center}
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   517
  \hfill
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   518
  \begin{tabular}{rcl}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   519
  \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   520
  @{text "ABS (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\\
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   521
  @{text "REP (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id :: \<sigma> \<Rightarrow> \<sigma>"}\smallskip\\
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   522
  \multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\ 
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   523
  @{text "ABS (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "REP (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> ABS (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\\
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   524
  @{text "REP (\<sigma>\<^isub>1 \<Rightarrow> \<sigma>\<^isub>2, \<tau>\<^isub>1 \<Rightarrow> \<tau>\<^isub>2)"} & $\dn$ & @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1) \<singlearr> REP (\<sigma>\<^isub>2, \<tau>\<^isub>2)"}\smallskip\\
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   525
  \multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\ 
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   526
  @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (ABS (\<sigma>s, \<tau>s))"}\\
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   527
  @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (REP (\<sigma>s, \<tau>s))"}\smallskip\\
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   528
  \multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors:}\\
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   529
  @{text "ABS (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "Abs_\<kappa>\<^isub>q \<circ> (MAP(\<rho>s \<kappa>) (ABS (\<sigma>s', \<tau>s)))"}\\
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   530
  @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "(MAP(\<rho>s \<kappa>) (REP (\<sigma>s', \<tau>s))) \<circ> Rep_\<kappa>\<^isub>q"}
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   531
  \end{tabular}\hfill\numbered{ABSREP}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   532
  \end{center}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   533
  %
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   534
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   535
  In the last two clauses we have that the type @{text "\<alpha>s
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   536
  \<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   537
  @{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   538
  list"}). The quotient construction ensures that the type variables in @{text
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   539
  "\<rho>s"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   540
  matchers for the @{text "\<alpha>s"} when matching @{text "\<rho>s \<kappa>"} against
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   541
  @{text "\<sigma>s \<kappa>"}.  The
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   542
  function @{text "MAP"} calculates an \emph{aggregate map-function} for a raw
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   543
  type as follows:
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   544
  %
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   545
  \begin{center}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   546
  \begin{tabular}{rcl}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   547
  @{text "MAP' (\<alpha>)"} & $\dn$ & @{text "a\<^sup>\<alpha>"}\\
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   548
  @{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id :: \<kappa> \<Rightarrow> \<kappa>"}\\
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   549
  @{text "MAP' (\<sigma>s \<kappa>)"} & $\dn$ & @{text "map_\<kappa> (MAP'(\<sigma>s))"}\smallskip\\
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   550
  @{text "MAP (\<sigma>)"} & $\dn$ & @{text "\<lambda>as. MAP'(\<sigma>)"}  
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   551
  \end{tabular}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   552
  \end{center}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   553
  %
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   554
  \noindent
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   555
  In this definition we rely on the fact that we can interpret type-variables @{text \<alpha>} as 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   556
  term variables @{text a}. In the last clause we build an abstraction over all
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   557
  term-variables of the map-function generated by the auxiliary function 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   558
  @{text "MAP'"}.
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   559
  The need for aggregate map-functions can be seen in cases where we build quotients, 
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   560
  say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types, say @{text "(\<alpha> list) \<times> \<beta>"}. 
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   561
  In this case @{text MAP} generates  the 
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   562
  aggregate map-function:
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   563
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   564
  @{text [display, indent=10] "\<lambda>a b. map_prod (map a) b"}
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   565
  
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   566
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   567
  which is essential in order to define the corresponding aggregate 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   568
  abstraction and representation functions.
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   569
  
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   570
  To see how these definitions pan out in practise, let us return to our
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   571
  example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   572
  @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   573
  fset"}. Feeding these types into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   574
  the abstraction function
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   575
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   576
  @{text [display, indent=10] "(map (map id \<circ> Rep_fset) \<circ> Rep_fset) \<singlearr> Abs_fset \<circ> map id"}
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   577
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   578
  \noindent
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   579
  In our implementation we further
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   580
  simplify this function by rewriting with the usual laws about @{text
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   581
  "map"}s and @{text "id"}, namely @{term "map id = id"} and @{text "f \<circ> id =
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   582
  id \<circ> f = f"}. This gives us the simpler abstraction function
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   583
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   584
  @{text [display, indent=10] "(map Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset"}
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   585
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   586
  \noindent
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   587
  which we can use for defining @{term "fconcat"} as follows
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   588
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   589
  @{text [display, indent=10] "\<Union> \<equiv> ((map Rep_fset \<circ> Rep_fset) \<singlearr> Abs_fset) flat"}
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   590
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   591
  \noindent
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   592
  Note that by using the operator @{text "\<singlearr>"} and special clauses
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   593
  for function types in \eqref{ABSREP}, we do not have to 
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   594
  distinguish between arguments and results, but can deal with them uniformly.
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   595
  Consequently, all definitions in the quotient package 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   596
  are of the general form
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   597
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   598
  @{text [display, indent=10] "c \<equiv> ABS (\<sigma>, \<tau>) t"}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   599
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   600
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   601
  where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   602
  type of the defined quotient constant @{text "c"}. This data can be easily
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   603
  generated from the declaration given by the user.
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   604
  To increase the confidence in this way of making definitions, we can prove 
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   605
  that the terms involved are all typable.
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   606
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   607
  \begin{lemma}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   608
  If @{text "ABS (\<sigma>, \<tau>)"} returns some abstraction function @{text "Abs"} 
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   609
  and @{text "REP (\<sigma>, \<tau>)"} some representation function @{text "Rep"}, 
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   610
  then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   611
  @{text "\<tau> \<Rightarrow> \<sigma>"}.
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   612
  \end{lemma}
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   613
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   614
  \begin{proof}
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   615
  By mutual induction and analysing the definitions of @{text "ABS"} and @{text "REP"}. 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   616
  The cases of equal types and function types are
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   617
  straightforward (the latter follows from @{text "\<singlearr>"} having the
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   618
  type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   619
  constructors we can observe that a map-function after applying the functions
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   620
  @{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}.  The
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   621
  interesting case is the one with unequal type constructors. Since we know
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   622
  the quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   623
  that @{text "Abs_\<kappa>\<^isub>q"} is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   624
  \<kappa>\<^isub>q"}. This type can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   625
  \<kappa>\<^isub>q"} where the type variables @{text "\<alpha>s"} are instantiated with the
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   626
  @{text "\<tau>s"}. The complete type can be calculated by observing that @{text
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   627
  "MAP (\<rho>s \<kappa>)"}, after applying the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it,
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   628
  returns a term of type @{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   629
  equivalent to @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}, which we just have to compose with
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   630
  @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"} according to the type of @{text "\<circ>"}.\qed
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   631
  \end{proof}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   632
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   633
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   634
  The reader should note that this lemma fails for the abstraction and representation 
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   635
  functions used in Homeier's quotient package.
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   636
*}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   637
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   638
section {* Respectfulness and Preservation \label{sec:resp} *}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   639
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   640
text {*
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   641
  The main point of the quotient package is to automatically ``lift'' theorems
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   642
  involving constants over the raw type to theorems involving constants over
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   643
  the quotient type. Before we can describe this lifting process, we need to impose 
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   644
  two restrictions in the form of proof obligations that arise during the
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   645
  lifting. The reason is that even if definitions for all raw constants 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   646
  can be given, \emph{not} all theorems can be lifted to the quotient type. Most 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   647
  notably is the bound variable function, that is the constant @{text bn}, defined 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   648
  for raw lambda-terms as follows
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   649
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   650
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2252
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   651
  @{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
Christian Urban <urbanc@in.tum.de>
parents: 2248
diff changeset
   652
  @{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\hspace{4mm}
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   653
  @{text "bn (\<lambda>x. t) \<equiv> {x} \<union> bn (t)"}
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   654
  \end{isabelle}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   655
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   656
  \noindent
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   657
  We can generate a definition for this constant using @{text ABS} and @{text REP}.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   658
  But this constant does \emph{not} respect @{text "\<alpha>"}-equivalence and 
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   659
  consequently no theorem involving this constant can be lifted to @{text
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   660
  "\<alpha>"}-equated lambda terms. Homeier formulates the restrictions in terms of
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   661
  the properties of \emph{respectfulness} and \emph{preservation}. We have
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   662
  to slightly extend Homeier's definitions in order to deal with quotient
2247
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   663
  compositions. 
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   664
084b2b7df98a some tuning and start work on section 4
Christian Urban <urbanc@in.tum.de>
parents: 2243
diff changeset
   665
  To formally define what respectfulness is, we have to first define 
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   666
  the notion of \emph{aggregate equivalence relations} using @{text REL}:
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   667
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   668
  \begin{center}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   669
  \hfill
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   670
  \begin{tabular}{rcl}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   671
  \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   672
  @{text "REL (\<sigma>, \<sigma>)"} & $\dn$ & @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}\smallskip\\
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   673
   \multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\ 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   674
  @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>)"} & $\dn$ & @{text "rel_\<kappa> (REL (\<sigma>s, \<tau>s))"}\smallskip\\
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   675
  \multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors:}\smallskip\\
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   676
  @{text "REL (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "rel_\<kappa>\<^isub>q (REL (\<sigma>s', \<tau>s))"}\\
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   677
  \end{tabular}\hfill\numbered{REL}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   678
  \end{center}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   679
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   680
  \noindent
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   681
  The @{text "\<sigma>s'"} in the last clause are calculated as in \eqref{ABSREP}:
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   682
  we know that type @{text "\<alpha>s \<kappa>\<^isub>q"} is the quotient of the raw type 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   683
  @{text "\<rho>s \<kappa>"}. The @{text "\<sigma>s'"} are determined by matching 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   684
  @{text "\<rho>s \<kappa>"} and @{text "\<sigma>s \<kappa>"}.
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   685
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   686
  Lets return to the lifting procedure of theorems. Assume we have a theorem
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   687
  that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   688
  lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   689
  constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   690
  we throw the following proof obligation
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   691
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   692
  @{text [display, indent=10] "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
2258
72ce58b76c3b finished preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2257
diff changeset
   693
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   694
  %%% PROBLEM I do not yet completely understand the 
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   695
  %%% form of respectfulness theorems
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   696
  %%%\noindent
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   697
  %%%if @ {text \<sigma>} and @ {text \<tau>} have no type variables. In case they have, then
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   698
  %%%the proof obligation is of the form
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   699
  %%% 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   700
  %%%@ {text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub>  implies  REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   701
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   702
  %%% ANSWER: The respectfulness theorems never have any quotient assumptions,
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   703
  %%% So the commited version is ok.
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   704
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   705
  \noindent
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   706
  %%%where @ {text "\<alpha>s"} are the type variables in @{text \<sigma>} and @{text \<tau>}.
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   707
  Homeier calls these proof obligations \emph{respectfulness
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   708
  theorems}. However, unlike his quotient package, we might have several
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   709
  respectfulness theorems for one constant---he has at most one.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   710
  The reason is that because of our quotient compositions, the types
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   711
  @{text \<sigma>} and @{text \<tau>} are not completely determined by the type of @{text "c\<^bsub>r\<^esub>"}.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   712
  And for every instantiation of the types, we might end up with a 
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   713
  corresponding respectfulness theorem.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   714
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   715
  Before lifting a theorem, we require the user to discharge
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   716
  them. And the point with @{text bn} is that the respectfulness theorem
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   717
  looks as follows
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   718
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   719
  @{text [display, indent=10] "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   720
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   721
  \noindent
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   722
  and the user cannot discharge it---because it is not true. To see this,
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   723
  we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun} 
2276
Christian Urban <urbanc@in.tum.de>
parents: 2275
diff changeset
   724
  using extensionally to obtain
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   725
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   726
  @{text [display, indent=10] "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 implies bn(t\<^isub>1) = bn(t\<^isub>2)"}
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   727
 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   728
  \noindent
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   729
  In contrast, if we lift a theorem about @{text "append"} to a theorem describing 
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   730
  the union of finite sets, then we need to discharge the proof obligation
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   731
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   732
  @{text [display, indent=10] "(\<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub> \<doublearr> \<approx>\<^bsub>list\<^esub>) append append"}
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   733
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   734
  \noindent
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   735
  To do so, we have to establish
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   736
  
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   737
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   738
  if @{thm (prem1) append_rsp_unfolded[of xs ys us vs, no_vars]} and
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   739
  @{thm (prem2) append_rsp_unfolded[of xs ys us vs, no_vars]}
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   740
  then @{thm (concl) append_rsp_unfolded[of xs ys us vs, no_vars]} 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   741
  \end{isabelle}
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   742
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   743
  \noindent
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   744
  which is straightforward given the definition shown in \eqref{listequiv}.
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   745
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   746
  The second restriction we have to impose arises from
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   747
  non-lifted polymorphic constants, which are instantiated to a
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   748
  type being quotient. For example, take the @{term "cons"} to 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   749
  add a pair of natural numbers to a list. The pair of natural numbers 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   750
  is to become an integer. But we still want to use @{text cons} for
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   751
  adding integers to lists---just with a different type. 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   752
  To be able to lift such theorems, we need a \emph{preservation theorem} 
2276
Christian Urban <urbanc@in.tum.de>
parents: 2275
diff changeset
   753
  for @{text cons}. Assuming we have a polymorphic raw constant 
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   754
  @{text "c\<^isub>r :: \<sigma>"} and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, 
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   755
  then a preservation theorem is as follows
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   756
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   757
  @{text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies  ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   758
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   759
  \noindent
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   760
  where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   761
  In case of @{text cons} (which has type @{text "\<alpha> \<Rightarrow> \<alpha> list \<Rightarrow> \<alpha> list"}) we have 
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   762
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   763
  @{text [display, indent=10] "(Rep ---> map Rep ---> map Abs) cons = cons"}
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   764
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   765
  \noindent
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   766
  under the assumption @{text "Quotient R Abs Rep"}. Interestingly, if we have
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   767
  an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   768
  with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   769
  then we need to show the corresponding preservation lemma.
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
   770
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   771
  @{thm [display, indent=10] insert_preserve2[no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   772
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   773
  %Given two quotients, one of which quotients a container, and the
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   774
  %other quotients the type in the container, we can write the
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   775
  %composition of those quotients. To compose two quotient theorems
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   776
  %we compose the relations with relation composition as defined above
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   777
  %and the abstraction and relation functions are the ones of the sub
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   778
  %quotients composed with the usual function composition.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   779
  %The @ {term "Rep"} and @ {term "Abs"} functions that we obtain agree
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   780
  %with the definition of aggregate Abs/Rep functions and the
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   781
  %relation is the same as the one given by aggregate relations.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   782
  %This becomes especially interesting
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   783
  %when we compose the quotient with itself, as there is no simple
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   784
  %intermediate step.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   785
  %
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   786
  %Lets take again the example of @ {term flat}. To be able to lift
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   787
  %theorems that talk about it we provide the composition quotient
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   788
  %theorem which allows quotienting inside the container:
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   789
  %
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   790
  %If @ {term R} is an equivalence relation and @ {term "Quotient R Abs Rep"}
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   791
  %then
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   792
  % 
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   793
  %@ {text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map Abs) (map Rep o rep_fset)"}
2275
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   794
  %%%
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   795
  %%%\noindent
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   796
  %%%this theorem will then instantiate the quotients needed in the
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   797
  %%%injection and cleaning proofs allowing the lifting procedure to
69b80ad616c5 finished section 4, but put some things I do not understand on comment
Christian Urban <urbanc@in.tum.de>
parents: 2274
diff changeset
   798
  %%%proceed in an unchanged way.
2192
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   799
*}
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   800
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
   801
section {* Lifting of Theorems\label{sec:lift} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   802
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   803
text {*
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   804
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   805
  The main benefit of a quotient package is to lift automatically theorems over raw
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   806
  types to theorems over quotient types. We will perform this lifting in
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   807
  three phases, called \emph{regularization},
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   808
  \emph{injection} and \emph{cleaning}.
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   809
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   810
  The purpose of regularization is to change the quantifiers and abstractions
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   811
  in a ``raw'' theorem to quantifiers over variables that respect the relation
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
   812
  (Definition \ref{def:respects} states what respects means). The purpose of injection is supposed to add @{term Rep}
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   813
  and @{term Abs} of appropriate types in front of constants and variables
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   814
  of the raw type so that they can be replaced by the ones that include the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   815
  quotient type. Cleaning rewrites the obtained injected theorem with
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   816
  preservation rules obtaining the desired goal theorem.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   817
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   818
  Most quotient packages take only an original theorem involving raw
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   819
  types and lift it. The procedure in our package takes both an
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   820
  original theorem involving raw types and a statement of the theorem
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   821
  that it is supposed to produce. To simplify the use of the quotient
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   822
  package we additionally provide an automated statement translation
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   823
  mechanism which can produce the latter automatically given a list of
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   824
  quotient types.  It is possible that a user wants to lift only some
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   825
  occurrences of a raw type. In this case the user specifies the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   826
  complete lifted goal instead of using the automated mechanism.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   827
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   828
  In the following we will first define the statement of the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   829
  regularized theorem based on the original theorem and the goal
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   830
  theorem. Then we define the statement of the injected theorem, based
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   831
  on the regularized theorem and the goal. We then show the 3 proofs,
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   832
  all three can be performed independently from each other.
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   833
2251
1a4fc8d3873f Qpaper / beginnig of sec5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2250
diff changeset
   834
  We define the function @{text REG}, which takes the statements
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   835
  of the raw theorem and the lifted theorem (both as terms) and
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   836
  returns the statement of the regularized version. The intuition
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   837
  behind this function is that it replaces quantifiers and
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   838
  abstractions involving raw types by bounded ones, and equalities
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   839
  involving raw types are replaced by appropriate aggregate
2251
1a4fc8d3873f Qpaper / beginnig of sec5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2250
diff changeset
   840
  equivalence relations. It is defined as follows:
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   841
2244
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   842
  \begin{center}
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   843
  \begin{longtable}{rcl}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   844
  \multicolumn{3}{@ {}l}{abstractions:}\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   845
  @{text "REG (\<lambda>x\<^sup>\<sigma>. t, \<lambda>x\<^sup>\<tau>. s)"} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   846
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   847
  @{text "\<lambda>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   848
  @{text "\<lambda>x\<^sup>\<sigma> \<in> Respects (REL (\<sigma>, \<tau>)). REG (t, s)"}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   849
  \end{cases}$\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   850
  \multicolumn{3}{@ {}l}{universal quantifiers:}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   851
  @{text "REG (\<forall>x\<^sup>\<sigma>. t, \<forall>x\<^sup>\<tau>. s)"} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   852
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   853
  @{text "\<forall>x\<^sup>\<sigma>. REG (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   854
  @{text "\<forall>x\<^sup>\<sigma> \<in> Respects (REL (\<sigma>, \<tau>)). REG (t, s)"}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   855
  \end{cases}$\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   856
  \multicolumn{3}{@ {}l}{equality:}\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   857
  @{text "REG (=\<^bsup>\<sigma>\<Rightarrow>\<sigma>\<Rightarrow>bool\<^esup>, =\<^bsup>\<tau>\<Rightarrow>\<tau>\<Rightarrow>bool\<^esup>)"} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   858
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   859
  @{text "="} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   860
  @{text "REL (\<sigma>, \<tau>)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   861
  \end{cases}$\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   862
  \multicolumn{3}{@ {}l}{applications, variables and constants:}\\
2244
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   863
  @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2)"} & $\dn$ & @{text "REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}\\
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   864
  @{text "REG (x\<^isub>1, x\<^isub>2)"} & $\dn$ & @{text "x\<^isub>1"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   865
  @{text "REG (c\<^isub>1, c\<^isub>2)"} & $\dn$ & @{text "c\<^isub>1"}\\[-5mm]
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   866
  \end{longtable}
2244
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   867
  \end{center}
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   868
  %
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   869
  \noindent
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   870
  In the above definition we omitted the cases for existential quantifiers
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   871
  and unique existential quantifiers, as they are very similar to the cases
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   872
  for the universal quantifier.
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   873
  Next we define the function @{text INJ} which takes the statement of
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   874
  the regularized theorems and the statement of the lifted theorem both as
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   875
  terms and returns the statement of the injected theorem:
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   876
2245
280b92df6a8b qpaper / INJ
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2244
diff changeset
   877
  \begin{center}
280b92df6a8b qpaper / INJ
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2244
diff changeset
   878
  \begin{tabular}{rcl}
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   879
  \multicolumn{3}{@ {\hspace{-4mm}}l}{abstractions:}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   880
  @{text "INJ (\<lambda>x. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   881
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   882
  @{text "\<lambda>x. INJ (t, s)"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   883
  @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x. INJ (t, s)))"}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   884
  \end{cases}$\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   885
  @{text "INJ (\<lambda>x \<in> R. t :: \<sigma>, \<lambda>x. s :: \<tau>) "} & $\dn$ 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   886
  & @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) (\<lambda>x \<in> R. INJ (t, s)))"}\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   887
  \multicolumn{3}{@ {\hspace{-4mm}}l}{universal quantifiers:}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   888
  @{text "INJ (\<forall> t, \<forall> s) "} & $\dn$ & @{text "\<forall> INJ (t, s)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   889
  @{text "INJ (\<forall> t \<in> R, \<forall> s) "} & $\dn$ & @{text "\<forall> INJ (t, s) \<in> R"}\smallskip\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   890
  \multicolumn{3}{@ {\hspace{-4mm}}l}{applications, variables and constants:}\smallskip\\
2245
280b92df6a8b qpaper / INJ
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2244
diff changeset
   891
  @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) "} & $\dn$ & @{text " INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}\\
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   892
  @{text "INJ (x\<^isub>1\<^sup>\<sigma>, x\<^isub>2\<^sup>\<tau>) "} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   893
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   894
  @{text "x\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   895
  @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) x\<^isub>1)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   896
  \end{cases}$\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   897
  @{text "INJ (c\<^isub>1\<^sup>\<sigma>, c\<^isub>2\<^sup>\<tau>) "} & $\dn$ & 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   898
  $\begin{cases}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   899
  @{text "c\<^isub>1"} \quad\mbox{provided @{text "\<sigma> = \<tau>"}}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   900
  @{text "REP (\<sigma>, \<tau>) (ABS (\<sigma>, \<tau>) c\<^isub>1)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   901
  \end{cases}$\\
2245
280b92df6a8b qpaper / INJ
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2244
diff changeset
   902
  \end{tabular}
280b92df6a8b qpaper / INJ
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2244
diff changeset
   903
  \end{center}
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   904
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   905
  \noindent where the cases for existential quantifiers and unique existential
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   906
  quantifiers have been omitted for clarity; are similar to universal quantifier.
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   907
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   908
  We can now define the subgoals that will imply the lifted theorem. Given
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   909
  the statement of the original theorem @{term t} and the statement of the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   910
  goal @{term g} the regularization subgoal is @{term "t \<longrightarrow> REG(t, g)"},
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   911
  the injection subgoal is @{term "REG(t, g) = INJ(REG(t, g), g)"} and the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   912
  cleaning subgoal is @{term "INJ(REG(t, g), g) = g"}. We will now describe
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   913
  the three tactics provided for these three subgoals.
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   914
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   915
  The injection and cleaning subgoals are always solved if the appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   916
  respectfulness and preservation theorems are given. It is not the case
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   917
  with regularization; sometimes a theorem given by the user does not
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   918
  imply a regularized version and a stronger one needs to be proved. This
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   919
  is outside of the scope of the quotient package, so such obligations are
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   920
  left to the user. Take a simple statement for integers @{text "0 \<noteq> 1"}.
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   921
  It does not follow from the fact that @{text "(0, 0) \<noteq> (1, 0)"} because
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   922
  of regularization. The raw theorem only shows that particular items in the
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   923
  equivalence classes are not equal. A more general statement saying that
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   924
  the classes are not equal is necessary.
2261
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2260
diff changeset
   925
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   926
  In the proof of the regularization subgoal we always start with an implication.
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   927
  Isabelle provides a set of \emph{mono} rules, that are used to split implications
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   928
  of similar statements into simpler implication subgoals. These are enhanced
2249
1476c26d4310 qpaper/unfold the ball_reg_right statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2246
diff changeset
   929
  with special quotient theorem in the regularization proof. Below we only show
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   930
  the versions for the universal quantifier. For the existential quantifier
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   931
  and abstraction they are analogous.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   932
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   933
  First, bounded universal quantifiers can be removed on the right:
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   934
2249
1476c26d4310 qpaper/unfold the ball_reg_right statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2246
diff changeset
   935
  @{thm [display, indent=10] ball_reg_right_unfolded[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   936
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   937
  They can be removed anywhere if the relation is an equivalence relation:
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   938
2265
9c44db3eef95 Remove only reference to 'equivp'.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2264
diff changeset
   939
  @{thm [display, indent=10] (concl) ball_reg_eqv[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   940
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   941
  And finally it can be removed anywhere if @{term R\<^isub>2} is an equivalence relation:
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   942
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   943
  @{thm [display, indent=10] (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   944
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   945
  The last theorem is new in comparison with Homeier's package. There the
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   946
  injection procedure would be used to prove goals with such shape, and there
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   947
  the equivalence assumption would be used. We use the above theorem directly
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   948
  also for composed relations where the range type is a type for which we know an
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   949
  equivalence theorem. This allows separating regularization from injection.
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   950
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   951
  The injection proof starts with an equality between the regularized theorem
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   952
  and the injected version. The proof again follows by the structure of the
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   953
  two terms, and is defined for a goal being a relation between these two terms.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   954
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   955
  \begin{itemize}
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
   956
  \item For two constants, an appropriate constant respectfulness assumption is used.
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   957
  \item For two variables, we use the assumptions proved in regularization.
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   958
  \item For two abstractions, they are eta-expanded and beta-reduced.
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   959
  \item For two applications, if the right side is an application of
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   960
    @{term Rep} to an @{term Abs} and @{term "Quotient R Rep Abs"} we
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   961
    can reduce the injected pair using the theorem:
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   962
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   963
    @{term [display, indent=10] "R x y \<longrightarrow> R x (Rep (Abs y))"}
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   964
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   965
    otherwise we introduce an appropriate relation between the subterms
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   966
    and continue with two subgoals using the lemma:
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   967
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   968
    @{text [display, indent=10] "(R\<^isub>1 \<doublearr> R\<^isub>2) f g \<longrightarrow> R\<^isub>1 x y \<longrightarrow> R\<^isub>2 (f x) (g y)"}
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   969
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   970
  \end{itemize}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   971
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   972
  The cleaning subgoal has been defined in such a way that
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   973
  establishing the goal theorem now consists only on rewriting the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   974
  injected theorem with the preservation theorems and quotient
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   975
  definitions. First for all lifted constants, their definitions
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   976
  are used to fold the @{term Rep} with the raw constant. Next for
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   977
  all lambda abstractions and quantifications the lambda and
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   978
  quantifier preservation theorems are used to replace the
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   979
  variables that include raw types with respects by quantification
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   980
  over variables that include quotient types. We show here only
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   981
  the lambda preservation theorem; assuming
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   982
  @{term "Quotient R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1"} and @{term "Quotient R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2"}
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   983
  hold, we have:
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   984
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   985
    @{thm [display, indent=10] (concl) lambda_prs[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   986
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   987
  \noindent
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   988
  holds. Next relations over lifted types are folded to equality.
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   989
  The following theorem has been shown in Homeier~\cite{Homeier05}:
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   990
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   991
    @{thm [display, indent=10] (concl) Quotient_rel_rep[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   992
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   993
  \noindent
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   994
  Finally the user given preservation theorems, that allow using
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   995
  higher level operations and containers of types being lifted.
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   996
  We show the preservation theorem for @{term map}. Again assuming
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
   997
  that @{term "Quotient R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1"} and @{term "Quotient R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2"} hold,
2271
c0c5bc4ee8cb qpaper/Rewrite section5
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2268
diff changeset
   998
  we have:
2212
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   999
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1000
  @{thm [display, indent=10] (concl) map_prs(1)[of R\<^isub>1 Abs\<^isub>1 Rep\<^isub>1 R\<^isub>2 Abs\<^isub>2 Rep\<^isub>2]}
2212
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
  1001
2246
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2245
diff changeset
  1002
  *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
  1003
2274
99cf23602d36 4 almost finished
Christian Urban <urbanc@in.tum.de>
parents: 2273
diff changeset
  1004
section {* Examples \label{sec:examples} *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
  1005
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1006
(* Mention why equivalence *)
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
  1007
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1008
text {*
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1009
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1010
  In this section we will show, a complete interaction with the quotient package
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1011
  for defining the type of integers by quotienting pairs of natural numbers and
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1012
  lifting theorems to integers. Our quotient package is fully compatible with
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1013
  Isabelle type classes, but for clarity we will not use them in this example.
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1014
  In a larger formalization of integers using the type class mechanism would
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1015
  provide many algebraic properties ``for free''.
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1016
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1017
  A user of our quotient package first needs to define a relation on
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1018
  the raw type, by which the quotienting will be performed. We give
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1019
  the same integer relation as the one presented in the introduction:
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1020
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1021
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1022
  \isacommand{fun}~~@{text "int_rel"}~~\isacommand{where}~~%
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1023
  @{text "(m :: nat, n) int_rel (p, q) = (m + q = n + p)"}
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1024
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1025
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1026
  \noindent
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1027
  Next the quotient type is defined. This leaves a proof obligation that the
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1028
  relation is an equivalence relation which is solved automatically using the
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1029
  definitions:
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1030
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1031
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
  1032
  \isacommand{quotient\_type}~~@{text "int"}~~\isacommand{=}~~@{text "(nat \<times> nat)"}~~\isacommand{/}~~@{text "int_rel"}
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1033
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1034
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
  1035
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1036
  The user can then specify the constants on the quotient type:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1037
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1038
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1039
  \begin{tabular}{@ {}l}
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1040
  \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1041
  \isacommand{fun}~~@{text "plus_raw"}~~\isacommand{where}~~%
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1042
  @{text "plus_raw (m, n) (p, q) = (m + p :: nat, n + q :: nat)"}\\
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1043
  \isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1044
  \isacommand{is}~~@{text "plus_raw"}\\
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1045
  \end{tabular}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1046
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1047
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1048
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1049
  Lets first take a simple theorem about addition on the raw level:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1050
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1051
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
  1052
  \isacommand{lemma}~~@{text "plus_zero_raw: int_rel (plus_raw (0, 0) x) x"}
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1053
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1054
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1055
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1056
  When the user tries to lift a theorem about integer addition, the respectfulness
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1057
  proof obligation is left, so let us prove it first:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1058
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1059
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1060
  \isacommand{lemma}~~@{text "[quot_respect]: 
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1061
  (int_rel \<doublearr> int_rel \<doublearr> int_rel) plus_raw plus_raw"}
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1062
  \end{isabelle}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1063
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1064
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1065
  Can be proved automatically by the system just by unfolding the definition
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1066
  of @{text "\<doublearr>"}.
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
  1067
  Now the user can either prove a lifted lemma explicitly:
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1068
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1069
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1070
  \isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting plus_zero_raw"}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1071
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1072
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1073
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1074
  Or in this simple case use the automated translation mechanism:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1075
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1076
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1077
  \isacommand{thm}~~@{text "plus_zero_raw[quot_lifted]"}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1078
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1079
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
  1080
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1081
  obtaining the same result.
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
  1082
*}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
  1083
2256
f5f21feaa168 some slight tuning of the preliminary section
Christian Urban <urbanc@in.tum.de>
parents: 2255
diff changeset
  1084
section {* Conclusion and Related Work\label{sec:conc}*}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
  1085
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
  1086
text {*
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
  1087
2267
3bcd715abd39 conclusion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2266
diff changeset
  1088
  The code of the quotient package and the examples described here are
3bcd715abd39 conclusion
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2266
diff changeset
  1089
  already included in the
2254
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2251
diff changeset
  1090
  standard distribution of Isabelle.\footnote{Available from
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1091
  \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} It is
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1092
  heavily used in the new version of Nominal Isabelle, which provides a convenient reasoning
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1093
  infrastructure for programming language calculi involving binders.  
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1094
  To achieve this, it builds types representing @{text \<alpha>}-equivalent terms.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1095
  Earlier
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1096
  versions of Nominal Isabelle have been used successfully in formalisations
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1097
  of an equivalence checking algorithm for LF \cite{UrbanCheneyBerghofer08},
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1098
  Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1099
  concurrency \cite{BengtsonParow09} and a strong normalisation result for
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1100
  cut-elimination in classical logic \cite{UrbanZhu08}.
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
  1101
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1102
  There is a wide range of existing of literature for dealing with
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1103
  quotients in theorem provers.
2273
d62c082cb56b cleaned up definitions
Christian Urban <urbanc@in.tum.de>
parents: 2272
diff changeset
  1104
  Slotosch~\cite{Slotosch97} implemented a mechanism that automatically
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1105
  defines quotient types for Isabelle/HOL. But he did not include theorem lifting.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1106
  Harrison's quotient package~\cite{harrison-thesis} is the first one that is
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1107
  able to automatically lift theorems, however only first-order theorems (that is theorems
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1108
  where abstractions, quantifiers and variables do not involve the quotient type). 
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1109
  There is also some work on quotient types in
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1110
  non-HOL based systems and logical frameworks, including theory interpretations
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1111
  in PVS~\cite{PVS:Interpretations}, new types in MetaPRL~\cite{Nogin02}, 
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1112
  and setoids in Coq \cite{ChicliPS02}.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1113
  Paulson showed a construction of quotients that does not require the
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1114
  Hilbert Choice operator, but also only first-order theorems to be lifted~\cite{Paulson06}.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1115
  The most related work to our package is the package for HOL4 by Homeier~\cite{Homeier05}.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1116
  He introduced most of the abstract notions about quotients and also deals with the
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1117
  lifting of higher-order theorems. However, he cannot deal with quotient compositions (needed
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1118
  for lifting theorems about @{text flat}. Also, a number of his definitions, like @{text ABS},
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1119
  @{text REP} and @{text INJ} etc only exist in ML-code. On the other hand, Homeier
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1120
  is able to deal with partial quotient constructions, which we have not implemented.
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1121
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1122
  One advantage of our package is that it is modular---in the sense that every step
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1123
  in the quotient construction can be done independently (see the criticism of Paulson
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1124
  about other quotient packages). This modularity is essential in the context of
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1125
  Isabelle, which supports type-classes and locales.
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1126
2278
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1127
  Another feature of our quotient package is that when lifting theorems, we can
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1128
  precisely specify what the lifted theorem should look like. This feature is
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1129
  necessary, for example, when lifting an inductive principle about two lists.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1130
  This principle has as the conclusion a predicate of the form @{text "P xs ys"},
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1131
  and we can precisely specify whether we want to quotient @{text "xs"} or @{text "ys"},
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1132
  or both. We found this feature very useful in the new version of Nominal 
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1133
  Isabelle.
337569f85398 conclusion done
Christian Urban <urbanc@in.tum.de>
parents: 2277
diff changeset
  1134
  \medskip
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1135
2263
d2ca79475103 qpaper/ackno
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2261
diff changeset
  1136
  \noindent
d2ca79475103 qpaper/ackno
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2261
diff changeset
  1137
  {\bf Acknowledgements:} We would like to thank Peter Homeier for the
2277
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
  1138
  discussions about his HOL4 quotient package and explaining to us
816204c76e90 Answer questions in comments
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2276
diff changeset
  1139
  some of its finer points in the implementation.
2263
d2ca79475103 qpaper/ackno
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2261
diff changeset
  1140
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1141
*}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1142
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1143
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
  1144
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1145
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1146
end
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
  1147
(*>*)