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