Quotient-Paper/Paper.thy
author Cezary Kaliszyk <kaliszyk@in.tum.de>
Mon, 14 Jun 2010 11:42:07 +0200
changeset 2244 e907165b953b
parent 2243 5e98b3f231a0
child 2245 280b92df6a8b
permissions -rw-r--r--
qpaper / REG
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
     1
(* How to change the notation for \<lbrakk> \<rbrakk> meta-level implications? *)
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
     2
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
theory Paper
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
     5
imports "Quotient"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
        "LaTeXsugar"
2186
762a739c9eb4 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2183
diff changeset
     7
        "../Nominal/FSet"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
begin
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     9
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    10
notation (latex output)
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    11
  rel_conj ("_ OOO _" [53, 53] 52) and
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    12
  "op -->" (infix "\<rightarrow>" 100) and
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    13
  "==>" (infix "\<Rightarrow>" 100) and
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    14
  fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    15
  fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    16
  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
    17
  fempty ("\<emptyset>") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    18
  funion ("_ \<union> _") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    19
  finsert ("{_} \<union> _") and 
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    20
  Cons ("_::_") and
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    21
  concat ("flat") and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    22
  fconcat ("\<Union>")
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    23
 
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    24
  
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    25
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    26
ML {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    27
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
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
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    37
setup {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    38
  Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    39
  Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    40
  Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    41
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
(*>*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    44
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
section {* Introduction *}
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    47
text {* 
2205
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    48
   \begin{flushright}
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    49
  {\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
    50
    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
    51
    Larry Paulson \cite{Paulson06}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    52
  \end{flushright}
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    53
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    54
  \noindent
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    55
  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
    56
  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
    57
  (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
    58
  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
    59
  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
    60
  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
    61
  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
    62
  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
    63
  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
    64
  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
    65
  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
    66
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    67
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    68
  @{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
    69
  \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
    70
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    71
  \noindent
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    72
  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
    73
  "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
    74
  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
    75
  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
    76
  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
    77
  "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
    78
  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
    79
  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
    80
  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
    81
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    82
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
    83
  @{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
    84
  \end{isabelle}
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    85
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    86
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    87
  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
    88
  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
    89
  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
    90
  @{text "\<union>"}, as list append.
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    91
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    92
  An area where quotients are ubiquitous is reasoning about programming language
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
    93
  calculi. A simple example is the lambda-calculus, whose raw terms are defined as
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    94
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
    95
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
    96
  @{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
    97
  \end{isabelle}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    98
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    99
  \noindent
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   100
  The problem with this definition arises when one attempts to
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   101
  prove formally, for example, the substitution lemma \cite{Barendregt81} by induction
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
   102
  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
   103
  \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
   104
  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
   105
  $\alpha$-equated lambda-terms, that means terms quotient according to
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   106
  $\alpha$-equivalence, then the reasoning infrastructure provided, 
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   107
  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
   108
  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
   109
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   110
  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
   111
  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
   112
  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
   113
  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
   114
  (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
   115
  usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.  
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   116
  It is feasible to to this work manually, if one has only a few quotient
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   117
  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
   118
  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
   119
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   120
  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
   121
  and automate the reasoning as much as possible. In the
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   122
  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
   123
  \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
   124
  \cite{Homeier05} implemented in HOL4.  The fundamental construction these
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   125
  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
   126
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   127
  \begin{center}
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   128
  \mbox{}\hspace{20mm}\begin{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   129
  %%\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
   130
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   131
  \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
   132
  \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
   133
  \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
   134
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   135
  \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
   136
  \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
   137
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   138
  \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
   139
  \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
   140
  \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
   141
    {\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
   142
  \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
   143
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   144
  \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
   145
  \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
   146
  \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
   147
  \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
   148
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   149
  \end{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   150
  \end{center}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   151
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   152
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   153
  The starting point is an existing type, to which we often refer as the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   154
  \emph{raw type}, over which an equivalence relation given by the user is
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   155
  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
   156
  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
   157
  \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
   158
  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
   159
  the functions @{text Abs} and @{text Rep} need to be derived from them. We
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   160
  will show the details later. } These functions relate elements in the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   161
  existing type to ones in the new type and vice versa; they can be uniquely
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   162
  identified by their type. For example for the integer quotient construction
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   163
  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
   164
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   165
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   166
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   167
  @{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
   168
  \end{isabelle}
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
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   171
  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
   172
  typing information is important. 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   173
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   174
  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
   175
  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
   176
  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
   177
  "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
   178
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   179
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   180
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   181
  @{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
   182
  \end{isabelle}
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
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   185
  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
   186
  @{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
   187
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   188
  @{text [display, indent=10] "add n m \<equiv> Abs_int (add_pair (Rep_int n) (Rep_int m))"}
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   189
  
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   190
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   191
  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
   192
  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
   193
  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
   194
  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
   195
  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
   196
  \emph{compositions} of quotients. Such compositions are needed for example
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   197
  in case of quotienting lists to obtain finite sets and the operator that 
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   198
  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
   199
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   200
  @{thm [display, indent=10] concat.simps(1) concat.simps(2)[no_vars]}
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
   201
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   202
  \noindent
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   203
  We expect that the corresponding operator on finite sets, written @{term "fconcat"},
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   204
  builds the union of finite sets of finite sets:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   205
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   206
  @{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
   207
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   208
  \noindent
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   209
  The quotient package should 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
   210
  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
   211
  that the method  used in the existing quotient
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   212
  packages of just taking the representation of the arguments and then take
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   213
  the abstraction of the result is \emph{not} enough. The reason is that case in case
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   214
  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
   215
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   216
  @{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
   217
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   218
  \noindent
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   219
  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
   220
  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
   221
  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
   222
  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
   223
  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
   224
  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
   225
  generate the following definition
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   226
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   227
  @{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
   228
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   229
  \noindent
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   230
  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
   231
  will present a formal definition of our aggregate abstraction and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   232
  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
   233
  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
   234
  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
   235
  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
   236
  of \emph{aggregate equivalence relation}, \emph{respectfulness} and \emph{preservation}
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   237
  from Homeier \cite{Homeier05}.
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   238
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   239
  We are also able to address the criticism by Paulson \cite{Paulson06} cited
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   240
  at the beginning of this section about having to collect theorems that are
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   241
  lifted from the raw level to the quotient level. Our quotient package is the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   242
  first one that is modular so that it allows to lift single theorems
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   243
  separately. This has the advantage for the user to develop a formal theory
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   244
  interactively an a natural progression. A pleasing result of the modularity
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   245
  is also that we are able to clearly specify what needs to be done in the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   246
  lifting process (this was only hinted at in \cite{Homeier05} and implemented
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   247
  as a ``rough recipe'' in ML-code).
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   248
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   249
  The paper is organised as follows: Section \ref{sec:prelims} presents briefly
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   250
  some necessary preliminaries; Section \ref{sec:type} presents the definitions 
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   251
  of quotient types and shows how definitions can be made over quotient types. \ldots
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
   252
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   253
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   254
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   255
section {* Preliminaries and General Quotient\label{sec:prelims} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   256
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   257
text {*
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   258
  In this section we present the definitions of a quotient that follow
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   259
  those by Homeier, the proofs can be found there.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   260
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   261
  \begin{definition}[Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   262
  A relation $R$ with an abstraction function $Abs$
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   263
  and a representation function $Rep$ is a \emph{quotient}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   264
  if and only if:
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   265
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   266
  \begin{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   267
  \item @{thm (rhs1) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   268
  \item @{thm (rhs2) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   269
  \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   270
  \end{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   271
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   272
  \end{definition}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   273
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   274
  \begin{definition}[Relation map and function map]\\
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   275
  @{thm fun_rel_def[of "R1" "R2", no_vars]}\\
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   276
  @{thm fun_map_def[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   277
  \end{definition}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   278
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   279
  The main theorems for building higher order quotients is:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   280
  \begin{lemma}[Function Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   281
  If @{thm (prem 1) fun_quotient[no_vars]} and @{thm (prem 2) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   282
  then @{thm (concl) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   283
  \end{lemma}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   284
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   285
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   286
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   287
subsection {* Higher Order Logic *}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   288
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   289
text {*
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   290
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   291
  Types:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   292
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   293
  @{text "\<sigma> ::="} & @{text "\<alpha>"} & \textrm{(type variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   294
      @{text "|"} & @{text "(\<sigma>,\<dots>,\<sigma>)\<kappa>"} & \textrm{(type construction)}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   295
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   296
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   297
  Terms:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   298
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   299
  @{text "t ::="} & @{text "x\<^isup>\<sigma>"} & \textrm{(variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   300
      @{text "|"} & @{text "c\<^isup>\<sigma>"} & \textrm{(constant)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   301
      @{text "|"} & @{text "t t"} & \textrm{(application)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   302
      @{text "|"} & @{text "\<lambda>x\<^isup>\<sigma>. t"} & \textrm{(abstraction)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   303
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   304
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   305
  {\it Say more about containers / maping functions }
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   306
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   307
  Such maps for most common types (list, pair, sum,
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   308
  option, \ldots) are described in Homeier, and we assume that @{text "map"}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   309
  is the function that returns a map for a given type.
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   310
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   311
  {\it say something about our use of @{text "\<sigma>s"}}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   312
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   313
*}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   314
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   315
section {* Quotient Types and Quotient Definitions\label{sec:type} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   316
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   317
text {*
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   318
  The first step in a quotient constructions 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
   319
  type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   320
  defined over a raw type, say @{text "\<sigma>"}. The type of this equivalence
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   321
  relation must be of type @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   322
  the declaration is therefore
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   323
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   324
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   325
  \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   326
  \end{isabelle}
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   327
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   328
  \noindent
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   329
  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
   330
  examples are
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   331
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   332
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   333
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   334
  \begin{tabular}{@ {}l}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   335
  \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
   336
  \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
   337
  \end{tabular}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   338
  \end{isabelle}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   339
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   340
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   341
  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
   342
  equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   343
  "\<approx>\<^bsub>list\<^esub>"} defined earlier in \eqref{natpairequiv} and
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   344
  \eqref{listequiv}, respectively.  Given this data, we declare internally 
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   345
  the quotient types as
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   346
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   347
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   348
  \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
   349
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   350
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   351
  \noindent
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   352
  where the right hand side is the (non-empty) set of equivalence classes of
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   353
  @{text "R"}. The restriction in this declaration is that the type variables
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   354
  in the raw type @{text "\<sigma>"} must be included in the type variables @{text
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   355
  "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will provide us with
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   356
  abstraction and representation functions having the type
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   357
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   358
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   359
  @{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
   360
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   361
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   362
  \noindent 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   363
  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
   364
  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
   365
  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
   366
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   367
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   368
  @{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
   369
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   370
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   371
  \noindent
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   372
  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
   373
  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
   374
  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
   375
  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
   376
  respectively.  Given that @{text "R"} is an equivalence relation, the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   377
  following property
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   378
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   379
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   380
  @{text [display, indent=10] "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   381
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   382
  \noindent
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   383
  holds (for the proof see \cite{Homeier05}).
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   384
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   385
  The next step in a quotient construction is to introduce new definitions
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   386
  involving the quotient type, which need to be defined in terms of concepts
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   387
  of the raw type (remember this is the only way how to extend HOL
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   388
  with new definitions). For the user visible is the declaration
2235
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   389
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   390
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   391
  \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
   392
  \end{isabelle}
ad725de6e39b more on the constant lifting section
Christian Urban <urbanc@in.tum.de>
parents: 2234
diff changeset
   393
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   394
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   395
  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
   396
  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
   397
  given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   398
  in places where a quotient and raw type are involved). Two concrete examples are
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   399
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   400
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   401
  \begin{tabular}{@ {}l}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   402
  \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
   403
  \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
   404
  \isacommand{is}~~@{text "flat"} 
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   405
  \end{tabular}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   406
  \end{isabelle}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   407
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   408
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   409
  The first one declares zero for integers and the second the operator for
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   410
  building unions of finite sets. 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   411
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   412
  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
   413
  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
   414
  quotient types involved. The data we rely on is the given quotient type
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   415
  @{text "\<tau>"} and the raw type @{text "\<sigma>"}.  They allow us to define aggregate
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   416
  abstraction and representation functions using the functions @{text "ABS (\<sigma>,
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   417
  \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses are given below. The idea behind
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   418
  them is to recursively descend into types @{text \<sigma>} and @{text \<tau>}, and generate the appropriate
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   419
  @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
2244
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   420
  we return just the identity whenever the types are equal. All clauses
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   421
  are as follows:
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   422
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   423
  \begin{center}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   424
  \begin{tabular}{rcl}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   425
  \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
   426
  @{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
   427
  @{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
   428
  \multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\ 
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   429
  @{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
   430
  @{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
   431
  \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
   432
  @{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
   433
  @{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
   434
  \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
   435
  @{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
   436
  @{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"}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   437
  \end{tabular}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   438
  \end{center}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   439
  %
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   440
  \noindent
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   441
  where in the last two clauses we have that the quotient 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
   442
  \<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
   443
  @{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
   444
  list"}). The quotient construction ensures that the type variables in @{text
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   445
  "\<rho>s"} must be amongst 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
   446
  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
   447
  @{text "\<sigma>s \<kappa>"}.  The
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   448
  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
   449
  type as follows:
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   450
  %
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   451
  \begin{center}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   452
  \begin{tabular}{rcl}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   453
  @{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
   454
  @{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
   455
  @{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
   456
  @{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
   457
  \end{tabular}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   458
  \end{center}
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   459
  %
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   460
  \noindent
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   461
  In this definition we abuse 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
   462
  term variables @{text a}. In the last clause we build an abstraction over all
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   463
  term-variables inside the aggregate map-function generated by the auxiliary function 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   464
  @{text "MAP'"}.
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   465
  The need of aggregate map-functions can be appreciated if we build quotients, 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   466
  say @{text "(\<alpha>, \<beta>) \<kappa>\<^isub>q"}, out of compound raw types of the form @{text "(\<alpha> list) \<times> \<beta>"}. 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   467
  In this case @{text MAP} generates the aggregate map-function:
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   468
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   469
  @{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
   470
  
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   471
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   472
  which we need to define the aggregate abstraction and representation functions.
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   473
  
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   474
  To se how these definitions pan out in practise, let us return to our
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   475
  example about @{term "concat"} and @{term "fconcat"}, where we have types
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   476
  @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   477
  fset"}. Feeding them into @{text ABS} gives us the abstraction function
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   478
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   479
  @{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
   480
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   481
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   482
  after some @{text "\<beta>"}-simplifications. In our implementation we further
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   483
  simplify this abstraction function employing the usual laws about @{text
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   484
  "map"}s and @{text "id"}, namely @{term "map id = id"} and @{text "f \<circ> id =
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   485
  id \<circ> f = f"}. This gives us the abstraction function
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   486
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   487
  @{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
   488
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   489
  \noindent
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   490
  which we can use for defining @{term "fconcat"} as follows
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   491
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   492
  @{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
   493
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   494
  \noindent
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   495
  Note that by using the operator @{text "\<singlearr>"} we do not have to 
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   496
  distinguish between arguments and results: the representation and abstraction
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   497
  functions are just inverses of each other, which we can combine using 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   498
  @{text "\<singlearr>"} to deal uniformly with arguments of functions and 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   499
  their result. As a result, all definitions in the quotient package 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   500
  are of the general form
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   501
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   502
  @{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
   503
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   504
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   505
  where @{text \<sigma>} is the type of the definiens @{text "t"} and @{text "\<tau>"} the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   506
  type of the defined quotient constant @{text "c"}. To ensure we
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   507
  obtained a correct definition, we can prove:
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   508
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   509
  \begin{lemma}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   510
  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
   511
  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
   512
  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
   513
  @{text "\<tau> \<Rightarrow> \<sigma>"}.
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   514
  \end{lemma}
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   515
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   516
  \begin{proof}
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   517
  By induction and analysing the definitions of @{text "ABS"}, @{text "REP"} 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   518
  and @{text "MAP"}. The cases of equal and function types are straightforward
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   519
  (the latter follows from @{text "\<singlearr>"} having the type
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   520
  @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). The case of equal
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   521
  type constructors follows by observing that a map-function after applying 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   522
  the functions @{text "ABS (\<sigma>s, \<tau>s)"} produce a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}.
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   523
  The interesting case is the one with unequal type constructors. Since we know the
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   524
  quotient is between @{text "\<alpha>s \<kappa>\<^isub>q"} and @{text "\<rho>s \<kappa>"}, we have that @{text "Abs_\<kappa>\<^isub>q"}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   525
  is of type @{text "\<rho>s \<kappa> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}, that can be more specialised to @{text "\<rho>s[\<tau>s] \<kappa> \<Rightarrow> \<tau>s \<kappa>\<^isub>q"}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   526
  where the type variables @{text "\<alpha>s"} are instantiated with @{text "\<tau>s"}. The 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   527
  complete type can be calculated by observing that @{text "MAP (\<rho>s \<kappa>)"} after applying
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   528
  the functions @{text "ABS (\<sigma>s', \<tau>s)"} to it, returns a term of type 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   529
  @{text "\<rho>s[\<sigma>s'] \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"}. This type is equivalent to  @{text "\<sigma>s \<kappa> \<Rightarrow> \<rho>s[\<tau>s] \<kappa>"} 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   530
  as desired.\qed
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   531
  \end{proof}
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   532
  
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   533
  \noindent
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   534
  The reader should note that this lemma fails for the abstraction and representation 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   535
  functions used, for example, in Homeier's quotient package.
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   536
*}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   537
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   538
section {* Respectfulness and Preservation *}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   539
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   540
text {*
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   541
  Before we can lift theorems involving the raw types to theorems over
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   542
  quotient types, we have to impose some restrictions. The reason is that not
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   543
  all theorems can be lifted. Homeier formulates these restrictions in terms
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   544
  of \emph{respectfullness} and \emph{preservation} of constants occuring in
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   545
  theorems.
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   546
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   547
  The respectfulness property for a constant states that it essentially 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   548
  respects the equivalence relation involved in the quotient. An example
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   549
  is the function returning bound variables of a lambda-term (see \eqref{lambda})
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   550
  and @{text "\<alpha>"}-equivalence. It will turn out that this function is not 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   551
  respectful. 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   552
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   553
  To state the respectfulness property we have to define \emph{aggregate equivalence 
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   554
  relations}.
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   555
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   556
  @{text [display] "GIVE DEFINITION HERE"}
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   557
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   558
  class returned by this constant depends only on the equivalence
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   559
  classes of the arguments applied to the constant. To automatically
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   560
  lift a theorem that talks about a raw constant, to a theorem about
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   561
  the quotient type a respectfulness theorem is required.
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   562
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   563
  A respectfulness condition for a constant can be expressed in
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   564
  terms of an aggregate relation between the constant and itself,
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   565
  for example the respectfullness for @{text "append"}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   566
  can be stated as:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   567
2238
8ddf1330f2ed completed proof and started section about respectfulness and preservation
Christian Urban <urbanc@in.tum.de>
parents: 2237
diff changeset
   568
  @{text [display, indent=10] "(R \<doublearr> R \<doublearr> R) append append"}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   569
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   570
  \noindent
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   571
  Which after unfolding the definition of @{term "op ===>"} is equivalent to:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   572
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   573
  @{thm [display, indent=10] append_rsp_unfolded[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   574
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   575
  \noindent An aggregate relation is defined in terms of relation
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   576
  composition, so we define it first:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   577
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   578
  \begin{definition}[Composition of Relations]
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   579
  @{abbrev "rel_conj R1 R2"} where @{text OO} is the predicate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   580
  composition @{thm pred_compI[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   581
  \end{definition}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   582
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   583
  The aggregate relation for an aggregate raw type and quotient type
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   584
  is defined as:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   585
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   586
  \begin{itemize}
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   587
  \item @{text "REL(\<alpha>\<^isub>1, \<alpha>\<^isub>2)"} = @{text "op ="}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   588
  \item @{text "REL(\<sigma>, \<sigma>)"}  =  @{text "op ="}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   589
  \item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>n))\<kappa>)"}  =  @{text "(rel \<kappa>) (REL(\<sigma>\<^isub>1,\<tau>\<^isub>1)) \<dots> (REL(\<sigma>\<^isub>n,\<tau>\<^isub>n))"}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   590
  \item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>\<^isub>1, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>m))\<kappa>\<^isub>2)"}  =  @{text "(rel \<kappa>\<^isub>1) (REL(\<rho>\<^isub>1,\<nu>\<^isub>1) \<dots> (REL(\<rho>\<^isub>p,\<nu>\<^isub>p) OOO Eqv_\<kappa>\<^isub>2"} provided @{text "\<eta> \<kappa>\<^isub>2 = (\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>p)\<kappa>\<^isub>1 \<and> \<exists>s. s(\<sigma>s\<kappa>\<^isub>1)=\<rho>s\<kappa>\<^isub>1 \<and> s(\<tau>s\<kappa>\<^isub>2)=\<nu>s\<kappa>\<^isub>2"}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   591
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   592
  \end{itemize}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   593
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   594
  Again, the last case is novel, so lets look at the example of
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   595
  respectfullness for @{term concat}. The statement according to
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   596
  the definition above is:
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   597
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   598
  @{thm [display, indent=10] concat_rsp[no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   599
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   600
  \noindent
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   601
  By unfolding the definition of relation composition and relation map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   602
  we can see the equivalent statement just using the primitive list
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   603
  equivalence relation:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   604
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   605
  @{thm [display, indent=10] concat_rsp_unfolded[of "a" "a'" "b'" "b", no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   606
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   607
  The statement reads that, for any lists of lists @{term a} and @{term b}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   608
  if there exist intermediate lists of lists @{term "a'"} and @{term "b'"}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   609
  such that each element of @{term a} is in the relation with an appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   610
  element of @{term a'}, @{term a'} is in relation with @{term b'} and each
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   611
  element of @{term b'} is in relation with the appropriate element of
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   612
  @{term b}.
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   613
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   614
*}
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   615
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   616
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   617
text {*
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   618
  Sometimes a non-lifted polymorphic constant is instantiated to a
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   619
  type being lifted. For example take the @{term "op #"} which inserts
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   620
  an element in a list of pairs of natural numbers. When the theorem
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   621
  is lifted, the pairs of natural numbers are to become integers, but
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   622
  the head constant is still supposed to be the head constant, just
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   623
  with a different type.  To be able to lift such theorems
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   624
  automatically, additional theorems provided by the user are
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   625
  necessary, we call these \emph{preservation} theorems following
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   626
  Homeier's naming.
2196
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   627
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   628
  To lift theorems that talk about insertion in lists of lifted types
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   629
  we need to know that for any quotient type with the abstraction and
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   630
  representation functions @{text "Abs"} and @{text Rep} we have:
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   631
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   632
  @{thm [display, indent=10] (concl) cons_prs[no_vars]}
2196
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   633
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   634
  This is not enough to lift theorems that talk about quotient compositions.
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   635
  For some constants (for example empty list) it is possible to show a
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   636
  general compositional theorem, but for @{term "op #"} it is necessary
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   637
  to show that it respects the particular quotient type:
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   638
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   639
  @{thm [display, indent=10] insert_preserve2[no_vars]}
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   640
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   641
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   642
subsection {* Composition of Quotient theorems *}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   643
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   644
text {*
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   645
  Given two quotients, one of which quotients a container, and the
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   646
  other quotients the type in the container, we can write the
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   647
  composition of those quotients. To compose two quotient theorems
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   648
  we compose the relations with relation composition as defined above
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   649
  and the abstraction and relation functions are the ones of the sub
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   650
  quotients composed with the usual function composition.
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   651
  The @{term "Rep"} and @{term "Abs"} functions that we obtain agree
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   652
  with the definition of aggregate Abs/Rep functions and the
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   653
  relation is the same as the one given by aggregate relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   654
  This becomes especially interesting
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   655
  when we compose the quotient with itself, as there is no simple
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   656
  intermediate step.
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   657
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   658
  Lets take again the example of @{term flat}. To be able to lift
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   659
  theorems that talk about it we provide the composition quotient
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   660
  theorems, which then lets us perform the lifting procedure in an
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   661
  unchanged way:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   662
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   663
  @{thm [display] quotient_compose_list[no_vars]}
2192
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   664
*}
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   665
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   666
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   667
section {* Lifting of Theorems *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   668
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   669
text {*
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   670
  The core of the quotient package takes an original theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   671
  talks about the raw types, and the statement of the theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   672
  it is supposed to produce. This is different from other existing
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   673
  quotient packages, where only the raw theorems were necessary.
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   674
  We notice that in some cases only some occurrences of the raw
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   675
  types need to be lifted. This is for example the case in the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   676
  new Nominal package, where a raw datatype that talks about
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   677
  pairs of natural numbers or strings (being lists of characters)
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   678
  should not be changed to a quotient datatype with constructors
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   679
  taking integers or finite sets of characters. To simplify the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   680
  use of the quotient package we additionally provide an automated
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   681
  statement translation mechanism that replaces occurrences of
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   682
  types that match given quotients by appropriate lifted types.
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   683
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   684
  Lifting the theorems is performed in three steps. In the following
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   685
  we call these steps \emph{regularization}, \emph{injection} and
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   686
  \emph{cleaning} following the names used in Homeier's HOL
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   687
  implementation.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   688
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   689
  We first define the statement of the regularized theorem based
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   690
  on the original theorem and the goal theorem. Then we define
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   691
  the statement of the injected theorem, based on the regularized
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   692
  theorem and the goal. We then show the 3 proofs, as all three
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   693
  can be performed independently from each other.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   694
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   695
*}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   696
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   697
subsection {* Regularization and Injection statements *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   698
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   699
text {*
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   700
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   701
  We first define the function @{text REG}, which takes the statements
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   702
  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
   703
  returns the statement of the regularized version. The intuition
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   704
  behind this function is that it replaces quantifiers and
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   705
  abstractions involving raw types by bounded ones, and equalities
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   706
  involving raw types are replaced by appropriate aggregate
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   707
  relations. It is defined as follows:
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   708
2244
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   709
  \begin{center}
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   710
  \begin{tabular}{rcl}
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   711
  \multicolumn{3}{@ {\hspace{-4mm}}l}{abstractions (with same types and different types):}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   712
  @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<sigma>. s)"} & $\dn$ & @{text "\<lambda>x : \<sigma>. REG (t, s)"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   713
  @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<tau>. s)"} & $\dn$ & @{text "\<lambda>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   714
  \multicolumn{3}{@ {\hspace{-4mm}}l}{quantification (over same types and different types):}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   715
  @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<sigma>. s)"} & $\dn$ & @{text "\<forall>x : \<sigma>. REG (t, s)"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   716
  @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<tau>. s)"} & $\dn$ & @{text "\<forall>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   717
  \multicolumn{3}{@ {\hspace{-4mm}}l}{equality (with same types and different types):}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   718
  @{text "REG ((op =) : \<sigma>, (op =) : \<sigma>)"} & $\dn$ & @{text "(op =) : \<sigma>"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   719
  @{text "REG ((op =) : \<sigma>, (op =) : \<tau>)"} & $\dn$ & @{text "REL (\<sigma>, \<tau>) : \<sigma>"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   720
  \multicolumn{3}{@ {\hspace{-4mm}}l}{application, variable, constant:}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   721
  @{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)"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   722
  @{text "REG (v\<^isub>1, v\<^isub>2)"} & $\dn$ & @{text "v\<^isub>1"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   723
  @{text "REG (c\<^isub>1, c\<^isub>2)"} & $\dn$ & @{text "c\<^isub>1"}\\
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   724
  \end{tabular}
e907165b953b qpaper / REG
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2243
diff changeset
   725
  \end{center}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   726
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   727
  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
   728
  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
   729
  for the universal quantifier.
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   730
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   731
  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
   732
  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
   733
  terms and returns the statement of the injected theorem:
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   734
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   735
  \begin{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   736
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<sigma>) = \<lambda>x. (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   737
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   738
  \item @{text "INJ ((\<lambda>x \<in> R. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x \<in> R. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   739
  \item @{text "INJ (\<forall> t, \<forall> s) = \<forall> (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   740
  \item @{text "INJ (\<forall> t \<in> R, \<forall> s) = \<forall> (INJ (t, s) \<in> R"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   741
  \item @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   742
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<sigma>) = v\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   743
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (v\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   744
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<sigma>) = c\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   745
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (c\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   746
  \end{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   747
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   748
  For existential quantifiers and unique existential quantifiers it is
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   749
  defined similarly to the universal one.
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   750
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   751
*}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   752
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   753
subsection {* Proof procedure *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   754
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   755
(* In the below the type-guiding 'QuotTrue' assumption is removed. We need it
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   756
   only for bound variables without types, while in the paper presentation
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   757
   variables are typed *)
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   758
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   759
text {*
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   760
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   761
  With the above definitions of @{text "REG"} and @{text "INJ"} we can show
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   762
  how the proof is performed. The first step is always the application of
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   763
  of the following lemma:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   764
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   765
  @{term [display, indent=10] "[|A; A --> B; B = C; C = D|] ==> D"}
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   766
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   767
  With @{text A} instantiated to the original raw theorem, 
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   768
       @{text B} instantiated to @{text "REG(A)"},
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   769
       @{text C} instantiated to @{text "INJ(REG(A))"},
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   770
   and @{text D} instantiated to the statement of the lifted theorem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   771
  The first assumption can be immediately discharged using the original
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   772
  theorem and the three left subgoals are exactly the subgoals of regularization,
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   773
  injection and cleaning. The three can be proved independently by the
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   774
  framework and in case there are non-solved subgoals they can be left
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   775
  to the user.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   776
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   777
  The injection and cleaning subgoals are always solved if the appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   778
  respectfulness and preservation theorems are given. It is not the case
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   779
  with regularization; sometimes a theorem given by the user does not
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   780
  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
   781
  is outside of the scope of the quotient package, so such obligations are
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   782
  left to the user. Take a simple statement for integers @{term "0 \<noteq> 1"}.
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   783
  It does not follow from the fact that @{term "\<not> (0, 0) = (1, 0)"} because
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   784
  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
   785
  equivalence classes are not equal. A more general statement saying that
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   786
  the classes are not equal is necessary.
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   787
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   788
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   789
subsection {* Proving Regularization *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   790
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   791
text {*
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   792
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   793
  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
   794
  of similar statements into simpler implication subgoals. These are enhanced
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   795
  with special quotient theorem in the regularization goal. Below we only show
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   796
  the versions for the universal quantifier. For the existential quantifier
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   797
  and abstraction they are analogous.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   798
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   799
  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
   800
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   801
  @{thm [display, indent=10] ball_reg_right[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   802
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   803
  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
   804
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   805
  @{thm [display, indent=10] ball_reg_eqv[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   806
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   807
  And finally it can be removed anywhere if @{term R2} is an equivalence relation, then:
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   808
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   809
  @{thm [display, indent=10] (concl) ball_reg_eqv_range[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   810
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   811
  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
   812
  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
   813
  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
   814
  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
   815
  equivalence theorem. This allows separating regularization from injection.
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   816
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   817
*}
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   818
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   819
(*
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   820
  @{thm bex_reg_eqv_range[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   821
  @{thm [display] bex_reg_left[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   822
  @{thm [display] bex1_bexeq_reg[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   823
  @{thm [display] bex_reg_eqv[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   824
  @{thm [display] babs_reg_eqv[no_vars]}
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   825
  @{thm [display] babs_simp[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   826
*)
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   827
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   828
subsection {* Injection *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   829
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   830
text {*
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   831
  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
   832
  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
   833
  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
   834
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   835
  \begin{itemize}
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   836
  \item For two constants, an appropriate constant respectfullness assumption is used.
2242
3f480e33d8df qpaper/various
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2241
diff changeset
   837
  \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
   838
  \item For two abstractions, they are eta-expanded and beta-reduced.
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   839
  \end{itemize}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   840
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   841
  Otherwise the two terms are applications. There are two cases: If there is a REP/ABS
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   842
  in the injected theorem we can use the theorem:
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   843
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   844
  @{thm [display, indent=10] rep_abs_rsp[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   845
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   846
  \noindent
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   847
  and continue the proof.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   848
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   849
  Otherwise we introduce an appropriate relation between the subterms and continue with
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   850
  two subgoals using the lemma:
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   851
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   852
  @{thm [display, indent=10] apply_rsp[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   853
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   854
*}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   855
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   856
subsection {* Cleaning *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   857
2212
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   858
text {*
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   859
  The @{text REG} and @{text INJ} functions have been defined in such a way
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   860
  that establishing the goal theorem now consists only on rewriting the
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   861
  injected theorem with the preservation theorems.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   862
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   863
  \begin{itemize}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   864
  \item First for lifted constants, their definitions are the preservation rules for
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   865
    them.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   866
  \item For lambda abstractions lambda preservation establishes
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   867
    the equality between the injected theorem and the goal. This allows both
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   868
    abstraction and quantification over lifted types.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   869
    @{thm [display] lambda_prs[no_vars]}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   870
  \item Relations over lifted types are folded with:
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   871
    @{thm [display] Quotient_rel_rep[no_vars]}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   872
  \item User given preservation theorems, that allow using higher level operations
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   873
    and containers of types being lifted. An example may be
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   874
    @{thm [display] map_prs(1)[of R1 Abs1 Rep1 R2 Abs2 Rep2,no_vars]}
2212
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   875
  \end{itemize}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   876
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   877
 Preservation of relations and user given constant preservation lemmas *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   878
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   879
section {* Examples *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   880
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   881
(* Mention why equivalence *)
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   882
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   883
text {*
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   884
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   885
  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
   886
  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
   887
  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
   888
  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
   889
  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
   890
  provide many algebraic properties ``for free''.
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   891
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   892
  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
   893
  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
   894
  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
   895
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   896
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
   897
  \isacommand{fun}~~@{text "int_rel"}~~\isacommand{where}~~@{text "(m \<Colon> nat, n) int_rel (p, q) = (m + q = n + p)"}
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   898
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   899
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   900
  \noindent
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   901
  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
   902
  relation is an equivalence relation which is solved automatically using the
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   903
  definitions:
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   904
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   905
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
   906
  \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
   907
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   908
2239
ff997de1bd73 qpaper/examples
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2238
diff changeset
   909
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   910
  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
   911
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   912
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   913
  \begin{tabular}{@ {}l}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   914
  \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   915
  \isacommand{fun}~~@{text "plus_raw"}~~\isacommand{where}~~@{text "plus_raw (m :: nat, n) (p, q) = (m + p, n + q)"}\\
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   916
  \isacommand{quotient\_definition}~~@{text "(op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int)"}~~\isacommand{is}~~@{text "plus_raw"}\\
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   917
  \end{tabular}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   918
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   919
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   920
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   921
  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
   922
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   923
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
   924
  \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
   925
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   926
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   927
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   928
  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
   929
  proof obligation is left, so let us prove it first:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   930
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   931
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
2241
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2240
diff changeset
   932
  \isacommand{lemma}~~@{text "[quot_respect]: (int_rel \<Longrightarrow> int_rel \<Longrightarrow> int_rel) plus_raw plus_raw"}
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   933
  \end{isabelle}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   934
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   935
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   936
  Can be proved automatically by the system just by unfolding the definition
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   937
  of @{text "op \<Longrightarrow>"}.
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   938
  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
   939
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   940
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   941
  \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
   942
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   943
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   944
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   945
  Or in this simple case use the automated translation mechanism:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   946
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   947
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   948
  \isacommand{thm}~~@{text "plus_zero_raw[quot_lifted]"}
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   949
  \end{isabelle}
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   950
2240
6c4b54482396 qpaper/more on example
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2239
diff changeset
   951
  \noindent
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   952
  obtaining the same result.
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   953
*}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   954
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   955
section {* Related Work *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   956
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   957
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   958
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   959
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   960
  \item Peter Homeier's package~\cite{Homeier05} (and related work from there)
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   961
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   962
  \item John Harrison's one~\cite{harrison-thesis} is the first one to lift theorems
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   963
    but only first order.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   964
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   965
  \item PVS~\cite{PVS:Interpretations}
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   966
  \item MetaPRL~\cite{Nogin02}
2243
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   967
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   968
%  \item Manually defined quotients in Isabelle/HOL Library (Markus's Quotient\_Type,
5e98b3f231a0 qpaper / minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2242
diff changeset
   969
%    Dixon's FSet, \ldots)
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   970
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   971
  \item Oscar Slotosch defines quotient-type automatically but no
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   972
    lifting~\cite{Slotosch97}.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   973
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   974
  \item PER. And how to avoid it.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   975
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   976
  \item Necessity of Hilbert Choice op and Larry's quotients~\cite{Paulson06}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   977
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   978
  \item Setoids in Coq and \cite{ChicliPS02}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   979
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   980
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   981
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   982
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   983
section {* Conclusion *}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   984
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   985
text {*
2237
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   986
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   987
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   988
  The code of the quotient package described here is already included in the
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   989
  standard distribution of Isabelle.\footnote{Avaiable from
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   990
  \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} It is
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   991
  heavily used in Nominal Isabelle, which provides a convenient reasoning
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   992
  infrastructure for programming language calculi involving binders.  Earlier
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   993
  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
   994
  of an equivalence checking algorithm for LF \cite{UrbanCheneyBerghofer08},
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   995
  Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   996
  concurrency \cite{BengtsonParow09} and a strong normalisation result for
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   997
  cut-elimination in classical logic \cite{UrbanZhu08}.
d1ab5d2d6926 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2236
diff changeset
   998
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   999
*}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1000
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1001
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1002
subsection {* Contributions *}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1003
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1004
text {*
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1005
  We present the detailed lifting procedure, which was not shown before.
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1006
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1007
  The quotient package presented in this paper has the following
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1008
  advantages over existing packages:
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1009
  \begin{itemize}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1010
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1011
  \item We define quotient composition, function map composition and
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1012
    relation map composition. This lets lifting polymorphic types with
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1013
    subtypes quotiented as well. We extend the notions of
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1014
    respectfulness and preservation to cope with quotient
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1015
    composition.
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1016
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1017
  \item We allow lifting only some occurrences of quotiented
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1018
    types. Rsp/Prs extended. (used in nominal)
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1019
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1020
  \item The quotient package is very modular. Definitions can be added
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1021
    separately, rsp and prs can be proved separately, Quotients and maps
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1022
    can be defined separately and theorems can
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1023
    be lifted on a need basis. (useful with type-classes).
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1024
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1025
  \item Can be used both manually (attribute, separate tactics,
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1026
    rsp/prs databases) and programatically (automated definition of
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1027
    lifted constants, the rsp proof obligations and theorem statement
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1028
    translation according to given quotients).
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1029
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1030
  \end{itemize}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1031
*}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1032
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
  1033
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
  1034
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1035
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
  1036
end
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
  1037
(*>*)