Quotient-Paper/Paper.thy
author Christian Urban <urbanc@in.tum.de>
Sun, 13 Jun 2010 17:01:15 +0200
changeset 2234 8035515bbbc6
parent 2233 22c6b6144abd
child 2235 ad725de6e39b
permissions -rw-r--r--
something about the quotient ype definitions
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
     1
(* How to change the notation for \<lbrakk> \<rbrakk> meta-level implications? *)
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
     2
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     3
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     4
theory Paper
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
     5
imports "Quotient"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     6
        "LaTeXsugar"
2186
762a739c9eb4 added FSet to the correct paper
Christian Urban <urbanc@in.tum.de>
parents: 2183
diff changeset
     7
        "../Nominal/FSet"
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
     8
begin
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
     9
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    10
notation (latex output)
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    11
  rel_conj ("_ OOO _" [53, 53] 52) and
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    12
  "op -->" (infix "\<rightarrow>" 100) and
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    13
  "==>" (infix "\<Rightarrow>" 100) and
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    14
  fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    15
  fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    16
  list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    17
  fempty ("\<emptyset>") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    18
  funion ("_ \<union> _") and
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    19
  finsert ("{_} \<union> _") and 
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    20
  Cons ("_::_") and
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
    21
  concat ("flat") and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    22
  fconcat ("\<Union>")
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    23
 
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    24
  
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
    25
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    26
ML {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    27
fun nth_conj n (_, r) = nth (HOLogic.dest_conj r) n;
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    28
fun style_lhs_rhs proj = Scan.succeed (fn ctxt => fn t =>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    29
  let
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    30
    val concl =
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    31
      Object_Logic.drop_judgment (ProofContext.theory_of ctxt) (Logic.strip_imp_concl t)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    32
  in
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    33
    case concl of (_ $ l $ r) => proj (l, r)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    34
    | _ => error ("Binary operator expected in term: " ^ Syntax.string_of_term ctxt concl)
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    35
  end);
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    36
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    37
setup {*
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    38
  Term_Style.setup "rhs1" (style_lhs_rhs (nth_conj 0)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    39
  Term_Style.setup "rhs2" (style_lhs_rhs (nth_conj 1)) #>
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    40
  Term_Style.setup "rhs3" (style_lhs_rhs (nth_conj 2))
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
    41
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    42
(*>*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    43
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
    44
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    45
section {* Introduction *}
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
    46
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    47
text {* 
2205
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    48
   \begin{flushright}
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    49
  {\em ``Not using a [quotient] package has its advantages: we do not have to\\ 
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    50
    collect all the theorems we shall ever want into one giant list;''}\\
2213
231a20534950 improved abstract, some tuning
Christian Urban <urbanc@in.tum.de>
parents: 2212
diff changeset
    51
    Larry Paulson \cite{Paulson06}
2205
69b4eb4b12c6 added larry's quote
Christian Urban <urbanc@in.tum.de>
parents: 2199
diff changeset
    52
  \end{flushright}\smallskip
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    53
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    54
  \noindent
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    55
  Isabelle is a popular generic theorem prover in which many logics can be
2214
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    56
  implemented. The most widely used one, however, is Higher-Order Logic
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    57
  (HOL). This logic consists of a small number of axioms and inference rules
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    58
  over a simply-typed term-language. Safe reasoning in HOL is ensured by two
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    59
  very restricted mechanisms for extending the logic: one is the definition of
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    60
  new constants in terms of existing ones; the other is the introduction of
02e03d4287ec a bit more in the introduction and abstract
Christian Urban <urbanc@in.tum.de>
parents: 2213
diff changeset
    61
  new types by identifying non-empty subsets in existing types. It is well
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
    62
  understood how to use both mechanisms for dealing with quotient
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    63
  constructions in HOL (see \cite{Homeier05,Paulson06}).  For example the
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    64
  integers in Isabelle/HOL are constructed by a quotient construction over the
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    65
  type @{typ "nat \<times> nat"} and the equivalence relation
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    66
2225
cbffed7d81bd Fix integer relation.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2224
diff changeset
    67
  @{text [display, indent=10] "(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"}
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    68
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
    69
  \noindent
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    70
  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
    71
  "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
    72
  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
    73
  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
    74
  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
    75
  "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
    76
  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
    77
  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
    78
  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
    79
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    80
  @{text [display, indent=10] "xs \<approx> ys \<equiv> (\<forall>x. memb x xs \<longleftrightarrow> memb x ys)"}
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    81
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
    82
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
    83
  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
    84
  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
    85
  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
    86
  @{text "\<union>"}, as list append.
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    87
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    88
  An area where quotients are ubiquitous is reasoning about programming language
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
    89
  calculi. A simple example is the lambda-calculus, whose raw terms are defined as
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    90
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    91
  @{text [display, indent=10] "t ::= x | t t | \<lambda>x.t"}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
    92
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
    93
  \noindent
2222
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
    94
  The problem with this definition arises when one attempts to
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
    95
  prove formally, for example, the substitution lemma \cite{Barendregt81} by induction
973649d612f8 more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2221
diff changeset
    96
  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
    97
  \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
    98
  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
    99
  $\alpha$-equated lambda-terms, that means terms quotient according to
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   100
  $\alpha$-equivalence, then the reasoning infrastructure provided, 
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   101
  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
   102
  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
   103
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   104
  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
   105
  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
   106
  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
   107
  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
   108
  (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
   109
  usually requires a \emph{lot} of tedious reasoning effort \cite{Paulson06}.  
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   110
  It is feasible to to this work manually, if one has only a few quotient
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   111
  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
   112
  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
   113
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   114
  The purpose of a \emph{quotient package} is to ease the lifting of theorems
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   115
  and automate the definitions and reasoning as much as possible. In the
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   116
  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
   117
  \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
   118
  \cite{Homeier05} implemented in HOL4.  The fundamental construction these
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   119
  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
   120
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   121
  \begin{center}
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   122
  \mbox{}\hspace{20mm}\begin{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   123
  %%\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
   124
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   125
  \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
   126
  \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
   127
  \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
   128
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   129
  \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
   130
  \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
   131
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   132
  \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
   133
  \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
   134
  \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
   135
    {\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
   136
  \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
   137
  
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   138
  \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
   139
  \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
   140
  \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
   141
  \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
   142
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   143
  \end{tikzpicture}
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   144
  \end{center}
2217
fc5bfd0cc1cd more on the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2215
diff changeset
   145
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   146
  \noindent
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   147
  The starting point is an existing type, often referred to as the \emph{raw
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   148
  type}, over which an equivalence relation given by the user is
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   149
  defined. With this input the package introduces a new type, to which we
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   150
  refer as the \emph{quotient type}. This type comes with an
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   151
  \emph{abstraction} and a \emph{representation} function, written @{text Abs}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   152
  and @{text Rep}.\footnote{Actually they need to be derived from slightly
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   153
  more basic functions. We will show the details later. } These functions
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   154
  relate elements in the existing type to ones in the new type and vice versa;
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   155
  they can be uniquely identified by their type. For example for the integer
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   156
  quotient construction the types of @{text Abs} and @{text Rep} are
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   157
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   158
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   159
  @{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
   160
  \end{isabelle}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   161
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   162
  \noindent
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   163
  However we often leave this typing information implicit for better
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   164
  readability, but sometimes write @{text Abs_int} and @{text Rep_int} if the
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   165
  typing information is important. Every abstraction and representation
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   166
  function stands for an isomorphism between the non-empty subset and elements
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   167
  in the new type. They are necessary for making definitions involving the new
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   168
  type. For example @{text "0"} and @{text "1"} of type @{typ int} can be
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   169
  defined as
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   170
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   171
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   172
  @{text "0 \<equiv> Abs (0, 0)"}\hspace{10mm}@{text "1 \<equiv> Abs (1, 0)"}
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   173
  \end{isabelle}
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   174
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   175
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   176
  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
   177
  @{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
   178
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   179
  @{text [display, indent=10] "add n m \<equiv> Abs (add_pair (Rep n) (Rep m))"}
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   180
  
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   181
  \noindent
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   182
  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
   183
  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
   184
  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
   185
  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
   186
  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
   187
  \emph{compositions} of quotients. Such compositions are needed for example
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   188
  in case of quotienting lists to obtain finite sets and the operator that 
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   189
  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
   190
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   191
  @{thm [display, indent=10] concat.simps(1) concat.simps(2)[no_vars]}
2183
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2182
diff changeset
   192
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   193
  \noindent
2221
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   194
  We expect that the corresponding operator on finite sets, written @{term "fconcat"},
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   195
  builds the union of finite sets of finite sets:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   196
2220
2c4c0d93daa6 more to the introduction of the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 2217
diff changeset
   197
  @{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
   198
2103
e08e3c29dbc0 a bit for the introduction of the q-paper
Christian Urban <urbanc@in.tum.de>
parents: 2102
diff changeset
   199
  \noindent
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   200
  The quotient package should provide us with a definition for @{text "\<Union>"} in
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   201
  terms of @{text flat}, @{text Rep} and @{text Abs} (the latter two having
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   202
  the type @{text "\<alpha> fset \<Rightarrow> \<alpha> list"} and @{text "\<alpha> list \<Rightarrow> \<alpha> fset"},
2226
36c9d9e658c7 some slight tuning of the intro
Christian Urban <urbanc@in.tum.de>
parents: 2225
diff changeset
   203
  respectively). The problem is that the method  used in the existing quotient
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   204
  packages of just taking the representation of the arguments and then take
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   205
  the abstraction of the result is \emph{not} enough. The reason is that case in case
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   206
  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
   207
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   208
  @{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
   209
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   210
  \noindent
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   211
  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
   212
  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
   213
  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
   214
  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
   215
  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
   216
  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
   217
  generate the following definition
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   218
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   219
  @{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
   220
e749cefbf66c even more on the qpaper (intro almost done)
Christian Urban <urbanc@in.tum.de>
parents: 2220
diff changeset
   221
  \noindent
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   222
  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
   223
  will present a formal definition of our aggregate abstraction and
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   224
  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
   225
  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
   226
  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
   227
  of the quotient constant. This means we also have to extend the notions
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   228
  of \emph{aggregate relation}, \emph{respectfulness} and \emph{preservation}
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   229
  from Homeier \cite{Homeier05}.
2223
c474186439bd more intro
Christian Urban <urbanc@in.tum.de>
parents: 2222
diff changeset
   230
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   231
  We will also address the criticism by Paulson \cite{Paulson06} cited at the
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   232
  beginning of this section about having to collect theorems that are lifted from the raw
2229
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2228
diff changeset
   233
  level to the quotient level. Our quotient package is the first one that is modular so that it
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   234
  allows to lift single theorems separately. This has the advantage for the
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   235
  user to develop a formal theory interactively an a natural progression. A
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   236
  pleasing result of the modularity is also that we are able to clearly
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   237
  specify what needs to be done in the lifting process (this was only hinted at in
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   238
  \cite{Homeier05} and implemented as a ``rough recipe'' in ML-code).
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   239
 
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   240
  The paper is organised as follows \ldots
2102
200954544cae added some of the quotient literature; a bit more to the qpaper
Christian Urban <urbanc@in.tum.de>
parents: 1994
diff changeset
   241
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   242
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   243
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   244
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   245
section {* Preliminaries and General Quotient\label{sec:prelims} *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   246
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   247
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   248
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   249
text {*
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   250
  In this section we present the definitions of a quotient that follow
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   251
  those by Homeier, the proofs can be found there.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   252
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   253
  \begin{definition}[Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   254
  A relation $R$ with an abstraction function $Abs$
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   255
  and a representation function $Rep$ is a \emph{quotient}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   256
  if and only if:
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   257
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   258
  \begin{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   259
  \item @{thm (rhs1) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   260
  \item @{thm (rhs2) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   261
  \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   262
  \end{enumerate}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   263
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   264
  \end{definition}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   265
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   266
  \begin{definition}[Relation map and function map]\\
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   267
  @{thm fun_rel_def[of "R1" "R2", no_vars]}\\
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   268
  @{thm fun_map_def[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   269
  \end{definition}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   270
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   271
  The main theorems for building higher order quotients is:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   272
  \begin{lemma}[Function Quotient]
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   273
  If @{thm (prem 1) fun_quotient[no_vars]} and @{thm (prem 2) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   274
  then @{thm (concl) fun_quotient[no_vars]}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   275
  \end{lemma}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   276
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   277
*}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   278
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   279
subsection {* Higher Order Logic *}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   280
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   281
text {*
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   282
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   283
  Types:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   284
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   285
  @{text "\<sigma> ::="} & @{text "\<alpha>"} & \textrm{(type variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   286
      @{text "|"} & @{text "(\<sigma>,\<dots>,\<sigma>)\<kappa>"} & \textrm{(type construction)}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   287
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   288
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   289
  Terms:
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   290
  \begin{eqnarray}\nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   291
  @{text "t ::="} & @{text "x\<^isup>\<sigma>"} & \textrm{(variable)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   292
      @{text "|"} & @{text "c\<^isup>\<sigma>"} & \textrm{(constant)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   293
      @{text "|"} & @{text "t t"} & \textrm{(application)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   294
      @{text "|"} & @{text "\<lambda>x\<^isup>\<sigma>. t"} & \textrm{(abstraction)} \\ \nonumber
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   295
  \end{eqnarray}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   296
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   297
  {\it Say more about containers / maping functions }
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   298
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   299
*}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   300
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   301
section {* Quotient Types and Lifting of Definitions *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   302
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   303
text {*
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   304
  The first step in a quotient constructions is to take a name for the new
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   305
  type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation defined over a
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   306
  raw type, say @{text "\<sigma>"}. The equivalence relation for the quotient
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   307
  construction, lets say @{text "R"}, must then be of type @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   308
  bool"}. Given this data, we can automatically declare the quotient type as
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   309
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   310
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   311
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   312
  \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
   313
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   314
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   315
  \noindent
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   316
  being the set of equivalence classes of @{text "R"}. The restriction in this declaration
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   317
  is that the type variables in the raw type @{text "\<sigma>"} must be included in the 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   318
  type variables @{text "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will provide us 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   319
  with abstraction and representation functions having the type
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   320
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   321
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   322
  @{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
   323
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   324
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   325
  \noindent
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   326
  relating the new quotient type and raw type. However, as Homeier noted, it is easier 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   327
  to work with the following derived definitions
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   328
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   329
  \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   330
  @{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
   331
  \end{isabelle}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   332
  
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   333
  \noindent
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   334
  on the expense of having to use Hilbert's choice operator @{text "\<epsilon>"}. With these
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   335
  definitions the abstraction and representation functions relate directly the 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   336
  quotient and raw types (their type is  @{text "\<sigma> \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"} and @{text "\<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma>"}, 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   337
  respectively). We can show that the property
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   338
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   339
  @{text [display, indent=10] "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   340
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   341
  \noindent
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   342
  holds (for the proof see \cite{Homeier05}).
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   343
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   344
  To define a constant on the lifted type, an aggregate abstraction
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   345
  function is applied to the raw constant. Below we describe the operation
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   346
  that generates
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   347
  an aggregate @{term "Abs"} or @{term "Rep"} function given the
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   348
  compound raw type and the compound quotient type.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   349
  This operation will also be used in translations of theorem statements
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   350
  and in the lifting procedure.
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   351
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   352
  The operation is additionally able to descend into types for which
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   353
  maps are known. Such maps for most common types (list, pair, sum,
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   354
  option, \ldots) are described in Homeier, and we assume that @{text "map"}
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   355
  is the function that returns a map for a given type. Then REP/ABS is defined
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   356
  as follows:
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   357
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   358
  {\it the first argument is the raw type and the second argument the quotient type}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   359
  %
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   360
  \begin{center}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   361
  \begin{longtable}{rcl}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   362
  \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   363
  @{text "ABS (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id"}\\
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   364
  @{text "REP (\<sigma>, \<sigma>)"} & $\dn$ & @{text "id"}\smallskip\\
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   365
  \multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\ 
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   366
  @{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
   367
  @{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
   368
  \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
   369
  @{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
   370
  @{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
   371
  \multicolumn{3}{@ {\hspace{-4mm}}l}{unequal type constructors:}\\
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   372
  @{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')))"}\\
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   373
  @{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"}\\
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   374
  \end{longtable}
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   375
  \end{center}
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   376
  %
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   377
  \noindent
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   378
  where in the last two clauses we have that the quotient type @{text "\<alpha>s \<kappa>\<^isub>q"} is a quotient of
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   379
  the raw type @{text "\<rho>s \<kappa>"} (for example @{text "\<alpha> fset"} and @{text "\<alpha> list"}). 
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   380
  The quotient construction ensures that the type variables in @{text "\<rho>s"} 
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   381
  must be amongst the @{text "\<alpha>s"}. Now the @{text "\<sigma>s'"} are given by the matchers 
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   382
  for the @{text "\<alpha>s"} 
2234
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   383
  when matching  @{text "\<sigma>s \<kappa>\<^isub>q"} against @{text "\<alpha>s \<kappa>\<^isub>q"}; similarly the @{text "\<tau>s'"} are given by the 
8035515bbbc6 something about the quotient ype definitions
Christian Urban <urbanc@in.tum.de>
parents: 2233
diff changeset
   384
  same type-variables when matching @{text "\<tau>s \<kappa>"} against @{text "\<rho>s \<kappa>"}.
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   385
  The function @{text "MAP"} calculates an \emph{aggregate map-function} for a type 
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   386
  as follows:
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   387
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   388
  \begin{center}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   389
  \begin{tabular}{rcl}
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   390
  @{text "MAP' (\<alpha>)"} & $\dn$ & @{text "a"}\\
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   391
  @{text "MAP' (\<kappa>)"} & $\dn$ & @{text "id"}\\
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   392
  @{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
   393
  @{text "MAP (\<sigma>)"} & $\dn$ & @{text "\<lambda>as. MAP'(\<sigma>)"}  
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   394
  \end{tabular}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   395
  \end{center}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   396
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   397
  \noindent
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   398
  In this definition we abuse the fact that we can interpret type-variables @{text \<alpha>} as 
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   399
  term variables @{text a}, and in the last clause build an abstraction over all
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   400
  term-variables inside the aggregate map-function generated by @{text "MAP'"}.
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   401
  This aggregate map-function is necessary if we build quotients, say @{text "(\<alpha>, \<beta>) foo"},
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   402
  out of compound raw types, say @{text "(\<alpha> list) \<times> \<beta>"}. In this case @{text MAP}
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   403
  generates the aggregate map-function:
2232
f49b5dfabd59 improved definition of ABS and REP
Christian Urban <urbanc@in.tum.de>
parents: 2231
diff changeset
   404
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   405
  @{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
   406
  
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   407
  \noindent
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   408
  Returning to our example about @{term "concat"} and @{term "fconcat"} which have the
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   409
  types @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}. Feeding this
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   410
  into @{text ABS} gives us the abstraction function:
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   411
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   412
  @{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
   413
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   414
  \noindent
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   415
  where we already performed some @{text "\<beta>"}-simplifications. In our implementation
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   416
  we further simplified this by noticing the usual laws about @{text "map"}s and @{text "id"},
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   417
  namely @{term "map id = id"} and 
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   418
  @{text "f \<circ> id = id \<circ> f = f"}. This gives us the simplified abstarction function
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   419
  
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   420
  @{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
   421
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   422
  \noindent
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   423
  which we can use for defining @{term "fconcat"} as follows
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   424
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   425
  @{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
   426
2195
0c1dcdefb515 Functionalized the ABS/REP definition.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2194
diff changeset
   427
  Apart from the last 2 points the definition is same as the one implemented in
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   428
  in Homeier's HOL package. Adding composition in last two cases is necessary
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   429
  for compositional quotients. We illustrate the different behaviour of the
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   430
  definition by showing the derived definition of @{term fconcat}:
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   431
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   432
  @{thm fconcat_def[no_vars]}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   433
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   434
  The aggregate @{term Abs} function takes a finite set of finite sets
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   435
  and applies @{term "map rep_fset"} composed with @{term rep_fset} to
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   436
  its input, obtaining a list of lists, passes the result to @{term concat}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   437
  obtaining a list and applies @{term abs_fset} obtaining the composed
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   438
  finite set.
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   439
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   440
  For every type map function we require the property that @{term "map id = id"}.
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   441
  With this we can compactify the term, removing the identity mappings,
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   442
  obtaining a definition that is the same as the one privided by Homeier
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   443
  in the cases where the internal type does not change.
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   444
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   445
  {\it we should be able to prove}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   446
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   447
  \begin{lemma}
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   448
  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
   449
  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
   450
  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
   451
  @{text "\<tau> \<Rightarrow> \<sigma>"}.
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   452
  \end{lemma}
2233
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   453
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   454
  This lemma fails in the over-simplified abstraction and representation function used
22c6b6144abd added some examples
Christian Urban <urbanc@in.tum.de>
parents: 2232
diff changeset
   455
  in Homeier's quotient package.
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   456
*}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   457
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   458
subsection {* Respectfulness *}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   459
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   460
text {*
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   461
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   462
  A respectfulness lemma for a constant states that the equivalence
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   463
  class returned by this constant depends only on the equivalence
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   464
  classes of the arguments applied to the constant. To automatically
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   465
  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
   466
  the quotient type a respectfulness theorem is required.
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   467
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   468
  A respectfulness condition for a constant can be expressed in
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   469
  terms of an aggregate relation between the constant and itself,
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   470
  for example the respectfullness for @{term "append"}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   471
  can be stated as:
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   472
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   473
  @{thm [display, indent=10] append_rsp[no_vars]}
2182
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2152
diff changeset
   474
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   475
  \noindent
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   476
  Which after unfolding the definition of @{term "op ===>"} is equivalent to:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   477
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   478
  @{thm [display, indent=10] append_rsp_unfolded[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   479
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   480
  \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
   481
  composition, so we define it first:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   482
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   483
  \begin{definition}[Composition of Relations]
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   484
  @{abbrev "rel_conj R1 R2"} where @{text OO} is the predicate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   485
  composition @{thm pred_compI[no_vars]}
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   486
  \end{definition}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   487
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   488
  The aggregate relation for an aggregate raw type and quotient type
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   489
  is defined as:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   490
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   491
  \begin{itemize}
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   492
  \item @{text "REL(\<alpha>\<^isub>1, \<alpha>\<^isub>2)"} = @{text "op ="}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   493
  \item @{text "REL(\<sigma>, \<sigma>)"}  =  @{text "op ="}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   494
  \item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>n))\<kappa>)"}  =  @{text "(rel \<kappa>) (REL(\<sigma>\<^isub>1,\<tau>\<^isub>1)) \<dots> (REL(\<sigma>\<^isub>n,\<tau>\<^isub>n))"}
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   495
  \item @{text "REL((\<sigma>\<^isub>1,\<dots>,\<sigma>\<^isub>n))\<kappa>\<^isub>1, (\<tau>\<^isub>1,\<dots>,\<tau>\<^isub>m))\<kappa>\<^isub>2)"}  =  @{text "(rel \<kappa>\<^isub>1) (REL(\<rho>\<^isub>1,\<nu>\<^isub>1) \<dots> (REL(\<rho>\<^isub>p,\<nu>\<^isub>p) OOO Eqv_\<kappa>\<^isub>2"} provided @{text "\<eta> \<kappa>\<^isub>2 = (\<alpha>\<^isub>1\<dots>\<alpha>\<^isub>p)\<kappa>\<^isub>1 \<and> \<exists>s. s(\<sigma>s\<kappa>\<^isub>1)=\<rho>s\<kappa>\<^isub>1 \<and> s(\<tau>s\<kappa>\<^isub>2)=\<nu>s\<kappa>\<^isub>2"}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   496
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   497
  \end{itemize}
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   498
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   499
  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
   500
  respectfullness for @{term concat}. The statement according to
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   501
  the definition above is:
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   502
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   503
  @{thm [display, indent=10] concat_rsp[no_vars]}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   504
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   505
  \noindent
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   506
  By unfolding the definition of relation composition and relation map
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   507
  we can see the equivalent statement just using the primitive list
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   508
  equivalence relation:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   509
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   510
  @{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
   511
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   512
  The statement reads that, for any lists of lists @{term a} and @{term b}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   513
  if there exist intermediate lists of lists @{term "a'"} and @{term "b'"}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   514
  such that each element of @{term a} is in the relation with an appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   515
  element of @{term a'}, @{term a'} is in relation with @{term b'} and each
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   516
  element of @{term b'} is in relation with the appropriate element of
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   517
  @{term b}.
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   518
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   519
*}
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   520
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   521
subsection {* Preservation *}
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   522
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   523
text {*
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   524
  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
   525
  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
   526
  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
   527
  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
   528
  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
   529
  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
   530
  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
   531
  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
   532
  Homeier's naming.
2196
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   533
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   534
  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
   535
  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
   536
  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
   537
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   538
  @{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
   539
74637f186af7 qpaper / a bit about prs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2195
diff changeset
   540
  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
   541
  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
   542
  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
   543
  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
   544
2228
a827d36fa467 qpaper / tuning in preservation and general display
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2227
diff changeset
   545
  @{thm [display, indent=10] insert_preserve2[no_vars]}
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   546
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   547
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   548
subsection {* Composition of Quotient theorems *}
2189
029bd37d010a qpaper..
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2188
diff changeset
   549
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   550
text {*
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   551
  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
   552
  other quotients the type in the container, we can write the
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   553
  composition of those quotients. To compose two quotient theorems
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   554
  we compose the relations with relation composition as defined above
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   555
  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
   556
  quotients composed with the usual function composition.
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   557
  The @{term "Rep"} and @{term "Abs"} functions that we obtain agree
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   558
  with the definition of aggregate Abs/Rep functions and the
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   559
  relation is the same as the one given by aggregate relations.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   560
  This becomes especially interesting
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   561
  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
   562
  intermediate step.
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   563
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   564
  Lets take again the example of @{term concat}. To be able to lift
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   565
  theorems that talk about it we provide the composition quotient
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   566
  theorems, which then lets us perform the lifting procedure in an
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   567
  unchanged way:
2188
57972032e20e qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2186
diff changeset
   568
2190
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2189
diff changeset
   569
  @{thm [display] quotient_compose_list[no_vars]}
2192
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   570
*}
87024a9a9d89 fixed compile error
Christian Urban <urbanc@in.tum.de>
parents: 2191
diff changeset
   571
2191
8fdfbec54229 qpaper / composition of quotients.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2190
diff changeset
   572
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   573
section {* Lifting of Theorems *}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   574
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   575
text {*
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   576
  The core of the quotient package takes an original theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   577
  talks about the raw types, and the statement of the theorem that
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   578
  it is supposed to produce. This is different from other existing
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   579
  quotient packages, where only the raw theorems were necessary.
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   580
  We notice that in some cases only some occurrences of the raw
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   581
  types need to be lifted. This is for example the case in the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   582
  new Nominal package, where a raw datatype that talks about
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   583
  pairs of natural numbers or strings (being lists of characters)
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   584
  should not be changed to a quotient datatype with constructors
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   585
  taking integers or finite sets of characters. To simplify the
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   586
  use of the quotient package we additionally provide an automated
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   587
  statement translation mechanism that replaces occurrences of
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   588
  types that match given quotients by appropriate lifted types.
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   589
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   590
  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
   591
  we call these steps \emph{regularization}, \emph{injection} and
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   592
  \emph{cleaning} following the names used in Homeier's HOL
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   593
  implementation.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   594
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   595
  We first define the statement of the regularized theorem based
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   596
  on the original theorem and the goal theorem. Then we define
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   597
  the statement of the injected theorem, based on the regularized
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   598
  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
   599
  can be performed independently from each other.
2193
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2192
diff changeset
   600
2194
a52499e125ce qpaper / lifting introduction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2193
diff changeset
   601
*}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   602
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   603
subsection {* Regularization and Injection statements *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   604
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   605
text {*
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   606
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   607
  We first define the function @{text REG}, which takes the statements
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   608
  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
   609
  returns the statement of the regularized version. The intuition
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   610
  behind this function is that it replaces quantifiers and
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   611
  abstractions involving raw types by bounded ones, and equalities
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   612
  involving raw types are replaced by appropriate aggregate
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   613
  relations. It is defined as follows:
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   614
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   615
  \begin{itemize}
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   616
  \item @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<sigma>. s) = \<lambda>x : \<sigma>. REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   617
  \item @{text "REG (\<lambda>x : \<sigma>. t, \<lambda>x : \<tau>. s) = \<lambda>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   618
  \item @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<sigma>. s) = \<forall>x : \<sigma>. REG (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   619
  \item @{text "REG (\<forall>x : \<sigma>. t, \<forall>x : \<tau>. s) = \<forall>x : \<sigma> \<in> Res (REL (\<sigma>, \<tau>)). REG (t, s)"}
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   620
  \item @{text "REG ((op =) : \<sigma>, (op =) : \<sigma>) = (op =) : \<sigma>"}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   621
  \item @{text "REG ((op =) : \<sigma>, (op =) : \<tau>) = REL (\<sigma>, \<tau>) : \<sigma>"}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   622
  \item @{text "REG (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = REG (t\<^isub>1, s\<^isub>1) REG (t\<^isub>2, s\<^isub>2)"}
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   623
  \item @{text "REG (v\<^isub>1, v\<^isub>2) = v\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   624
  \item @{text "REG (c\<^isub>1, c\<^isub>2) = c\<^isub>1"}
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   625
  \end{itemize}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   626
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   627
  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
   628
  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
   629
  for the universal quantifier.
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   630
2207
ea7c3f21d6df Qpaper/more.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2206
diff changeset
   631
  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
   632
  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
   633
  terms and returns the statement of the injected theorem:
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   634
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   635
  \begin{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   636
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<sigma>) = \<lambda>x. (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   637
  \item @{text "INJ ((\<lambda>x. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   638
  \item @{text "INJ ((\<lambda>x \<in> R. t) : \<sigma>, (\<lambda>x. s) : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (\<lambda>x \<in> R. (INJ (t, s))))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   639
  \item @{text "INJ (\<forall> t, \<forall> s) = \<forall> (INJ (t, s)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   640
  \item @{text "INJ (\<forall> t \<in> R, \<forall> s) = \<forall> (INJ (t, s) \<in> R"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   641
  \item @{text "INJ (t\<^isub>1 t\<^isub>2, s\<^isub>1 s\<^isub>2) = INJ (t\<^isub>1, s\<^isub>1) INJ (t\<^isub>2, s\<^isub>2)"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   642
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<sigma>) = v\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   643
  \item @{text "INJ (v\<^isub>1 : \<sigma>, v\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (v\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   644
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<sigma>) = c\<^isub>1"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   645
  \item @{text "INJ (c\<^isub>1 : \<sigma>, c\<^isub>2 : \<tau>) = REP(\<sigma>,\<tau>) (ABS (\<sigma>,\<tau>) (c\<^isub>1))"}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   646
  \end{itemize}
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   647
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   648
  For existential quantifiers and unique existential quantifiers it is
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   649
  defined similarly to the universal one.
2198
8fe1a706ade7 qpaper / injection statement
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2197
diff changeset
   650
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   651
*}
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   652
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   653
subsection {* Proof procedure *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   654
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   655
(* In the below the type-guiding 'QuotTrue' assumption is removed; since we
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   656
   present in a paper a version with typed-variables it is not necessary *)
2197
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   657
3a6afcb187ec qpaper / regularize
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2196
diff changeset
   658
text {*
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   659
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   660
  With the above definitions of @{text "REG"} and @{text "INJ"} we can show
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   661
  how the proof is performed. The first step is always the application of
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   662
  of the following lemma:
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   663
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   664
  @{term [display, indent=10] "[|A; A --> B; B = C; C = D|] ==> D"}
2208
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   665
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   666
  With @{text A} instantiated to the original raw theorem, 
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   667
       @{text B} instantiated to @{text "REG(A)"},
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   668
       @{text C} instantiated to @{text "INJ(REG(A))"},
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   669
   and @{text D} instantiated to the statement of the lifted theorem.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   670
  The first assumption can be immediately discharged using the original
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   671
  theorem and the three left subgoals are exactly the subgoals of regularization,
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   672
  injection and cleaning. The three can be proved independently by the
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   673
  framework and in case there are non-solved subgoals they can be left
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   674
  to the user.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   675
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   676
  The injection and cleaning subgoals are always solved if the appropriate
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   677
  respectfulness and preservation theorems are given. It is not the case
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   678
  with regularization; sometimes a theorem given by the user does not
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   679
  imply a regularized version and a stronger one needs to be proved. This
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   680
  is outside of the scope of the quotient package, so the user is then left
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   681
  with such obligations. As an example lets see the simplest possible
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   682
  non-liftable theorem for integers: When we want to prove @{term "0 \<noteq> 1"}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   683
  on integers the fact that @{term "\<not> (0, 0) = (1, 0)"} is not enough. It
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   684
  only shows that particular items in the equivalence classes are not equal,
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   685
  a more general statement saying that the classes are not equal is necessary.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   686
*}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   687
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   688
subsection {* Proving Regularization *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   689
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2207
diff changeset
   690
text {*
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   691
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   692
  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
   693
  of similar statements into simpler implication subgoals. These are enhanced
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   694
  with special quotient theorem in the regularization goal. Below we only show
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   695
  the versions for the universal quantifier. For the existential quantifier
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   696
  and abstraction they are analogous with some symmetry.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   697
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   698
  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
   699
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   700
  @{thm [display, indent=10] ball_reg_right[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   701
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   702
  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
   703
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   704
  @{thm [display, indent=10] ball_reg_eqv[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   705
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   706
  And finally it can be removed anywhere if @{term R2} is an equivalence relation, then:
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   707
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   708
  @{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
   709
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   710
  The last theorem is new in comparison with Homeier's package, there the
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   711
  injection procedure would be used to prove goals with such shape, and there
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   712
  the equivalence assumption would be useful. We use it directly also for
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   713
  composed relations where the range type is a type for which we know an
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   714
  equivalence theorem. This allows separating regularization from injection.
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   715
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   716
*}
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   717
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   718
(*
2231
01d08af79f01 qpaper.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2230
diff changeset
   719
  @{thm bex_reg_eqv_range[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   720
  @{thm [display] bex_reg_left[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   721
  @{thm [display] bex1_bexeq_reg[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   722
  @{thm [display] bex_reg_eqv[no_vars]}
2209
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   723
  @{thm [display] babs_reg_eqv[no_vars]}
5952b0f28261 Qpaper/regularization proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2208
diff changeset
   724
  @{thm [display] babs_simp[no_vars]}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   725
*)
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   726
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   727
subsection {* Injection *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   728
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   729
text {*
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   730
  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
   731
  and the injected version. The proof again follows by the structure of the
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   732
  two term, and is defined for a goal being a relation between the two terms.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   733
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   734
  \begin{itemize}
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   735
  \item For two constants, an appropriate constant respectfullness assumption is used.
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   736
  \item For two variables, the regularization assumptions state that they are related.
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   737
  \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
   738
  \end{itemize}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   739
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   740
  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
   741
  in the injected theorem we can use the theorem:
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   742
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   743
  @{thm [display] rep_abs_rsp[no_vars]}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   744
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   745
  and continue the proof.
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   746
2211
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   747
  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
   748
  two subgoals using the lemma:
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   749
9d0673c319d1 qpaper / injection proof.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2210
diff changeset
   750
  @{thm [display] apply_rsp[no_vars]}
2199
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   751
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   752
*}
6ce64fb5cbd9 qpaper / lemmas used in proofs
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2198
diff changeset
   753
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   754
subsection {* Cleaning *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   755
2212
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   756
text {*
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   757
  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
   758
  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
   759
  injected theorem with the preservation theorems.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   760
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   761
  \begin{itemize}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   762
  \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
   763
    them.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   764
  \item For lambda abstractions lambda preservation establishes
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   765
    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
   766
    abstraction and quantification over lifted types.
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   767
    @{thm [display] lambda_prs[no_vars]}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   768
  \item Relations over lifted types are folded with:
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   769
    @{thm [display] Quotient_rel_rep[no_vars]}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   770
  \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
   771
    and containers of types being lifted. An example may be
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   772
    @{thm [display] map_prs(1)[no_vars]}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   773
  \end{itemize}
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   774
79cebcc230d6 Qpaper / minor on cleaning
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2211
diff changeset
   775
 Preservation of relations and user given constant preservation lemmas *}
1994
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   776
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   777
section {* Examples *}
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1978
diff changeset
   778
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   779
(* Mention why equivalence *)
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   780
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   781
text {*
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   782
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   783
  A user of our quotient package first needs to define an equivalence relation:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   784
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   785
  @{text "fun \<approx> where (x, y) \<approx> (u, v) = (x + v = u + y)"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   786
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   787
  Then the user defines a quotient type:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   788
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   789
  @{text "quotient_type int = (nat \<times> nat) / \<approx>"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   790
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   791
  Which leaves a proof obligation that the relation is an equivalence relation,
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   792
  that can be solved with the automatic tactic with two definitions.
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   793
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   794
  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
   795
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   796
  @{text "quotient_definition 0 \<Colon> int is (0\<Colon>nat, 0\<Colon>nat)"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   797
  @{text "fun plus_raw where plus_raw (x, y) (u, v) = (x + u, y + v)"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   798
  @{text "quotient_definition (op +) \<Colon> (int \<Rightarrow> int \<Rightarrow> int) is plus_raw"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   799
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   800
  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
   801
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   802
  @{text "lemma plus_zero_raw: plus_raw (0, 0) i \<approx> i"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   803
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   804
  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
   805
  proof obligation is left, so let us prove it first:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   806
  
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   807
  @{text "lemma (op \<approx> \<Longrightarrow> op \<approx> \<Longrightarrow> op \<approx>) plus_raw plus_raw"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   808
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   809
  Can be proved automatically by the system just by unfolding the definition
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   810
  of @{term "op \<Longrightarrow>"}.
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   811
2230
fec38b7ceeb3 some spelling
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2229
diff changeset
   812
  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
   813
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   814
  @{text "lemma 0 + i = i by lifting plus_zero_raw"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   815
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   816
  Or in this simple case use the automated translation mechanism:
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   817
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   818
  @{text "thm plus_zero_raw[quot_lifted]"}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   819
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   820
  obtaining the same result.
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   821
*}
2206
2d6cada7d5e0 Qpaper/Minor
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2205
diff changeset
   822
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   823
section {* Related Work *}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   824
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   825
text {*
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   826
  \begin{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   827
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   828
  \item Peter Homeier's package~\cite{Homeier05} (and related work from there)
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   829
  \item John Harrison's one~\cite{harrison-thesis} is the first one to lift theorems
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   830
    but only first order.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   831
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   832
  \item PVS~\cite{PVS:Interpretations}
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   833
  \item MetaPRL~\cite{Nogin02}
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   834
  \item Manually defined quotients in Isabelle/HOL Library (Markus's Quotient\_Type,
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   835
    Dixon's FSet, \ldots)
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   836
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   837
  \item Oscar Slotosch defines quotient-type automatically but no
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   838
    lifting~\cite{Slotosch97}.
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   839
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   840
  \item PER. And how to avoid it.
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   841
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   842
  \item Necessity of Hilbert Choice op and Larry's quotients~\cite{Paulson06}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   843
2152
d7d4491535a9 starting bibliography
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2103
diff changeset
   844
  \item Setoids in Coq and \cite{ChicliPS02}
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   845
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   846
  \end{itemize}
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   847
*}
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   848
2210
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   849
section {* Conclusion *}
6aaec9dd0c62 qpaper / example interaction
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 2209
diff changeset
   850
2224
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   851
text {*
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   852
  The package is part of the standard distribution of Isabelle.
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   853
*}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   854
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   855
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   856
subsection {* Contributions *}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   857
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   858
text {*
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   859
  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
   860
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   861
  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
   862
  advantages over existing packages:
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   863
  \begin{itemize}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   864
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   865
  \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
   866
    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
   867
    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
   868
    respectfulness and preservation to cope with quotient
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   869
    composition.
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   870
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   871
  \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
   872
    types. Rsp/Prs extended. (used in nominal)
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   873
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   874
  \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
   875
    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
   876
    can be defined separately and theorems can
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   877
    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
   878
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   879
  \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
   880
    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
   881
    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
   882
    translation according to given quotients).
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   883
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   884
  \end{itemize}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   885
*}
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   886
f5b6f9d8a882 completed the intro (except minor things)
Christian Urban <urbanc@in.tum.de>
parents: 2223
diff changeset
   887
2227
42d576c54704 polishing of ABS/REP
Christian Urban <urbanc@in.tum.de>
parents: 2226
diff changeset
   888
1975
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   889
(*<*)
b1281a0051ae added stub for quotient paper; call with isabelle make qpaper
Christian Urban <urbanc@in.tum.de>
parents:
diff changeset
   890
end
1978
8feedc0d4ea8 quotient paper
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents: 1975
diff changeset
   891
(*>*)