Quotient-Paper/Paper.thy
changeset 2287 adb5b1349280
parent 2286 e7bc2ae30faf
child 2319 7c8783d2dcd0
equal deleted inserted replaced
2286:e7bc2ae30faf 2287:adb5b1349280
     8 (****
     8 (****
     9 
     9 
    10 ** things to do for the next version
    10 ** things to do for the next version
    11 *
    11 *
    12 * - what are quot_thms?
    12 * - what are quot_thms?
    13 * - what do all preservation theorems look like
    13 * - what do all preservation theorems look like,
       
    14     in particular preservation for quotient
       
    15     compositions
    14 *)
    16 *)
    15 
    17 
    16 notation (latex output)
    18 notation (latex output)
    17   rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
    19   rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
    18   pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
    20   pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
   296   @{text "\<sigma>s"} to stand for collections of type variables and types,
   298   @{text "\<sigma>s"} to stand for collections of type variables and types,
   297   respectively.  The type of a term is often made explicit by writing @{text
   299   respectively.  The type of a term is often made explicit by writing @{text
   298   "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
   300   "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function
   299   type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
   301   type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains many primitive and defined
   300   constants; a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
   302   constants; a primitive constant is equality, with type @{text "= :: \<sigma> \<Rightarrow> \<sigma> \<Rightarrow>
   301   bool"}, and the identity function with type @{text "id :: \<sigma> => \<sigma>"} is
   303   bool"}, and the identity function with type @{text "id :: \<sigma> \<Rightarrow> \<sigma>"} is
   302   defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}).
   304   defined as @{text "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}).
   303 
   305 
   304   An important point to note is that theorems in HOL can be seen as a subset
   306   An important point to note is that theorems in HOL can be seen as a subset
   305   of terms that are constructed specially (namely through axioms and prove
   307   of terms that are constructed specially (namely through axioms and proof
   306   rules). As a result we are able to define automatic proof
   308   rules). As a result we are able to define automatic proof
   307   procedures showing that one theorem implies another by decomposing the term
   309   procedures showing that one theorem implies another by decomposing the term
   308   underlying the first theorem.
   310   underlying the first theorem.
   309 
   311 
   310   Like Homeier, our work relies on map-functions defined for every type
   312   Like Homeier, our work relies on map-functions defined for every type
   402   First, a general quotient theorem, like the one given in Proposition \ref{funquot},
   404   First, a general quotient theorem, like the one given in Proposition \ref{funquot},
   403   cannot be stated inside HOL, because of the restriction on types.
   405   cannot be stated inside HOL, because of the restriction on types.
   404   Second, even if we were able to state such a quotient theorem, it
   406   Second, even if we were able to state such a quotient theorem, it
   405   would not be true in general. However, we can prove specific instances of a
   407   would not be true in general. However, we can prove specific instances of a
   406   quotient theorem for composing particular quotient relations.
   408   quotient theorem for composing particular quotient relations.
   407   To lift @{term flat} the quotient theorem for composing @{text "\<approx>\<^bsub>list\<^esub>"}
   409   For example, to lift theorems involving @{term flat} the quotient theorem for 
   408   is necessary:
   410   composing @{text "\<approx>\<^bsub>list\<^esub>"} will be necessary: given @{term "Quotient R Abs Rep"} 
   409 % It says when lists are being quotiented to finite sets,
   411   with @{text R} being an equivalence relation, then
   410 % the contents of the lists can be quotiented as well
   412 
   411  If @{term R} is an equivalence relation and @{term "Quotient R Abs Rep"} then
   413   @{text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map Abs) (map Rep \<circ> rep_fset)"}
   412 
       
   413   @{text [display, indent=10] "Quotient (list_rel R \<circ>\<circ>\<circ> \<approx>\<^bsub>list\<^esub>) (abs_fset \<circ> map Abs) (map Rep o rep_fset)"}
       
   414 
   414 
   415   \vspace{-.5mm}
   415   \vspace{-.5mm}
   416 
       
   417 
       
   418 *}
   416 *}
   419 
   417 
   420 section {* Quotient Types and Quotient Definitions\label{sec:type} *}
   418 section {* Quotient Types and Quotient Definitions\label{sec:type} *}
   421 
   419 
   422 text {*
   420 text {*
   482   respectively.  Given that @{text "R"} is an equivalence relation, the
   480   respectively.  Given that @{text "R"} is an equivalence relation, the
   483   following property holds  for every quotient type 
   481   following property holds  for every quotient type 
   484   (for the proof see \cite{Homeier05}).
   482   (for the proof see \cite{Homeier05}).
   485 
   483 
   486   \begin{proposition}
   484   \begin{proposition}
   487   @{text "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}
   485   @{text "Quotient R Abs_\<kappa>\<^isub>q Rep_\<kappa>\<^isub>q"}.
   488   \end{proposition}
   486   \end{proposition}
   489 
   487 
   490   The next step in a quotient construction is to introduce definitions of new constants
   488   The next step in a quotient construction is to introduce definitions of new constants
   491   involving the quotient type. These definitions need to be given in terms of concepts
   489   involving the quotient type. These definitions need to be given in terms of concepts
   492   of the raw type (remember this is the only way how to extend HOL
   490   of the raw type (remember this is the only way how to extend HOL
   659   involving constants over the raw type to theorems involving constants over
   657   involving constants over the raw type to theorems involving constants over
   660   the quotient type. Before we can describe this lifting process, we need to impose 
   658   the quotient type. Before we can describe this lifting process, we need to impose 
   661   two restrictions in the form of proof obligations that arise during the
   659   two restrictions in the form of proof obligations that arise during the
   662   lifting. The reason is that even if definitions for all raw constants 
   660   lifting. The reason is that even if definitions for all raw constants 
   663   can be given, \emph{not} all theorems can be lifted to the quotient type. Most 
   661   can be given, \emph{not} all theorems can be lifted to the quotient type. Most 
   664   notably is the bound variable function, that is the constant @{text bn}, defined 
   662   notable is the bound variable function, that is the constant @{text bn}, defined 
   665   for raw lambda-terms as follows
   663   for raw lambda-terms as follows
   666 
   664 
   667   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   665   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   668   @{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
   666   @{text "bn (x) \<equiv> \<emptyset>"}\hspace{4mm}
   669   @{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\hspace{4mm}
   667   @{text "bn (t\<^isub>1 t\<^isub>2) \<equiv> bn (t\<^isub>1) \<union> bn (t\<^isub>2)"}\hspace{4mm}
   678   the properties of \emph{respectfulness} and \emph{preservation}. We have
   676   the properties of \emph{respectfulness} and \emph{preservation}. We have
   679   to slightly extend Homeier's definitions in order to deal with quotient
   677   to slightly extend Homeier's definitions in order to deal with quotient
   680   compositions. 
   678   compositions. 
   681 
   679 
   682   To formally define what respectfulness is, we have to first define 
   680   To formally define what respectfulness is, we have to first define 
   683   the notion of \emph{aggregate equivalence relations} using @{text REL}:
   681   the notion of \emph{aggregate equivalence relations} using the function @{text REL}:
   684 
   682 
   685   \begin{center}
   683   \begin{center}
   686   \hfill
   684   \hfill
   687   \begin{tabular}{rcl}
   685   \begin{tabular}{rcl}
   688   \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
   686   \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
   702 
   700 
   703   Lets return to the lifting procedure of theorems. Assume we have a theorem
   701   Lets return to the lifting procedure of theorems. Assume we have a theorem
   704   that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
   702   that contains the raw constant @{text "c\<^isub>r :: \<sigma>"} and which we want to
   705   lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
   703   lift to a theorem where @{text "c\<^isub>r"} is replaced by the corresponding
   706   constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation 
   704   constant @{text "c\<^isub>q :: \<tau>"} defined over a quotient type. In this situation 
   707   we throw the following proof obligation
   705   we generate the following proof obligation
   708 
   706 
   709   @{text [display, indent=10] "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
   707   @{text [display, indent=10] "REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
   710 
   708 
   711   %%% PROBLEM I do not yet completely understand the 
   709   \noindent
   712   %%% form of respectfulness theorems
       
   713   %%%\noindent
       
   714   %%%if @ {text \<sigma>} and @ {text \<tau>} have no type variables. In case they have, then
       
   715   %%%the proof obligation is of the form
       
   716   %%% 
       
   717   %%%@ {text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub>  implies  REL (\<sigma>, \<tau>) c\<^isub>r c\<^isub>r"}
       
   718 
       
   719   %%% ANSWER: The respectfulness theorems never have any quotient assumptions,
       
   720   %%% So the commited version is ok.
       
   721 
       
   722   \noindent
       
   723   %%%where @ {text "\<alpha>s"} are the type variables in @{text \<sigma>} and @{text \<tau>}.
       
   724   Homeier calls these proof obligations \emph{respectfulness
   710   Homeier calls these proof obligations \emph{respectfulness
   725   theorems}. However, unlike his quotient package, we might have several
   711   theorems}. However, unlike his quotient package, we might have several
   726   respectfulness theorems for one constant---he has at most one.
   712   respectfulness theorems for one constant---he has at most one.
   727   The reason is that because of our quotient compositions, the types
   713   The reason is that because of our quotient compositions, the types
   728   @{text \<sigma>} and @{text \<tau>} are not completely determined by the type of @{text "c\<^bsub>r\<^esub>"}.
   714   @{text \<sigma>} and @{text \<tau>} are not completely determined by the type of @{text "c\<^bsub>r\<^esub>"}.
   729   And for every instantiation of the types, we might end up with a 
   715   And for every instantiation of the types, we might end up with a 
   730   corresponding respectfulness theorem.
   716   corresponding respectfulness theorem.
   731 
   717 
   732   Before lifting a theorem, we require the user to discharge
   718   Before lifting a theorem, we require the user to discharge
   733   them. And the point with @{text bn} is that the respectfulness theorem
   719   respectfulness proof obligations. And the point with @{text bn} is that the respectfulness theorem
   734   looks as follows
   720   looks as follows
   735 
   721 
   736   @{text [display, indent=10] "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
   722   @{text [display, indent=10] "(\<approx>\<^isub>\<alpha> \<doublearr> =) bn bn"}
   737 
   723 
   738   \noindent
   724   \noindent
   739   and the user cannot discharge it---because it is not true. To see this,
   725   and the user cannot discharge it: because it is not true. To see this,
   740   we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun} 
   726   we can just unfold the definition of @{text "\<doublearr>"} \eqref{relfun} 
   741   using extensionally to obtain
   727   using extensionally to obtain
   742 
   728 
   743   @{text [display, indent=10] "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 implies bn(t\<^isub>1) = bn(t\<^isub>2)"}
   729   @{text [display, indent=10] "\<forall>t\<^isub>1 t\<^isub>2. if t\<^isub>1 \<approx>\<^isub>\<alpha> t\<^isub>2 implies bn(t\<^isub>1) = bn(t\<^isub>2)"}
   744  
   730  
   750 
   736 
   751   \noindent
   737   \noindent
   752   To do so, we have to establish
   738   To do so, we have to establish
   753   
   739   
   754   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   740   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   755   if @{thm (prem1) append_rsp_unfolded[of xs ys us vs, no_vars]} and
   741   if @{text "xs \<approx>\<^bsub>list\<^esub> ys"} and  @{text "us \<approx>\<^bsub>list\<^esub> vs"}
   756   @{thm (prem2) append_rsp_unfolded[of xs ys us vs, no_vars]}
   742   then @{text "xs @ us \<approx>\<^bsub>list\<^esub> ys @ vs"}
   757   then @{thm (concl) append_rsp_unfolded[of xs ys us vs, no_vars]} 
       
   758   \end{isabelle}
   743   \end{isabelle}
   759 
   744 
   760   \noindent
   745   \noindent
   761   which is straightforward given the definition shown in \eqref{listequiv}.
   746   which is straightforward given the definition shown in \eqref{listequiv}.
   762 
   747 
   763   The second restriction we have to impose arises from
   748   The second restriction we have to impose arises from
   764   non-lifted polymorphic constants, which are instantiated to a
   749   non-lifted polymorphic constants, which are instantiated to a
   765   type being quotient. For example, take the @{term "cons"} to 
   750   type being quotient. For example, take the @{term "cons"}-constructor to 
   766   add a pair of natural numbers to a list. The pair of natural numbers 
   751   add a pair of natural numbers to a list, whereby teh pair of natural numbers 
   767   is to become an integer. But we still want to use @{text cons} for
   752   is to become an integer in te quotient construction. The point is that we 
       
   753   still want to use @{text cons} for
   768   adding integers to lists---just with a different type. 
   754   adding integers to lists---just with a different type. 
   769   To be able to lift such theorems, we need a \emph{preservation theorem} 
   755   To be able to lift such theorems, we need a \emph{preservation property} 
   770   for @{text cons}. Assuming we have a polymorphic raw constant 
   756   for @{text cons}. Assuming we have a polymorphic raw constant 
   771   @{text "c\<^isub>r :: \<sigma>"} and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, 
   757   @{text "c\<^isub>r :: \<sigma>"} and a corresponding quotient constant @{text "c\<^isub>q :: \<tau>"}, 
   772   then a preservation theorem is as follows
   758   then a preservation property is as follows
   773 
   759 
   774   @{text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies  ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
   760   @{text [display, indent=10] "Quotient R\<^bsub>\<alpha>s\<^esub> Abs\<^bsub>\<alpha>s\<^esub> Rep\<^bsub>\<alpha>s\<^esub> implies  ABS (\<sigma>, \<tau>) c\<^isub>r = c\<^isub>r"}
   775 
   761 
   776   \noindent
   762   \noindent
   777   where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
   763   where the @{text "\<alpha>s"} stand for the type variables in the type of @{text "c\<^isub>r"}.
   781 
   767 
   782   \noindent
   768   \noindent
   783   under the assumption @{text "Quotient R Abs Rep"}. Interestingly, if we have
   769   under the assumption @{text "Quotient R Abs Rep"}. Interestingly, if we have
   784   an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
   770   an instance of @{text cons} where the type variable @{text \<alpha>} is instantiated
   785   with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
   771   with @{text "nat \<times> nat"} and we also quotient this type to yield integers,
   786   then we need to show the corresponding preservation lemma.
   772   then we need to show the corresponding preservation property.
   787 
   773 
   788   %%%@ {thm [display, indent=10] insert_preserve2[no_vars]}
   774   %%%@ {thm [display, indent=10] insert_preserve2[no_vars]}
   789 
   775 
   790   %Given two quotients, one of which quotients a container, and the
   776   %Given two quotients, one of which quotients a container, and the
   791   %other quotients the type in the container, we can write the
   777   %other quotients the type in the container, we can write the
   820 text {*
   806 text {*
   821 
   807 
   822   The main benefit of a quotient package is to lift automatically theorems over raw
   808   The main benefit of a quotient package is to lift automatically theorems over raw
   823   types to theorems over quotient types. We will perform this lifting in
   809   types to theorems over quotient types. We will perform this lifting in
   824   three phases, called \emph{regularization},
   810   three phases, called \emph{regularization},
   825   \emph{injection} and \emph{cleaning}.
   811   \emph{injection} and \emph{cleaning} according to procedures in Homeier's ML-code.
   826 
   812 
   827   The purpose of regularization is to change the quantifiers and abstractions
   813   The purpose of regularization is to change the quantifiers and abstractions
   828   in a ``raw'' theorem to quantifiers over variables that respect the relation
   814   in a ``raw'' theorem to quantifiers over variables that respect the relation
   829   (Definition \ref{def:respects} states what respects means). The purpose of injection is to add @{term Rep}
   815   (Definition \ref{def:respects} states what respects means). The purpose of injection is to add @{term Rep}
   830   and @{term Abs} of appropriate types in front of constants and variables
   816   and @{term Abs} of appropriate types in front of constants and variables
   840   3.) & @{text "inj_thm \<longleftrightarrow> quot_thm"}\\
   826   3.) & @{text "inj_thm \<longleftrightarrow> quot_thm"}\\
   841   \end{tabular}
   827   \end{tabular}
   842   \end{center}
   828   \end{center}
   843 
   829 
   844   \noindent
   830   \noindent
       
   831   which means the raw theorem implies the quotient theorem.
   845   In contrast to other quotient packages, our package requires
   832   In contrast to other quotient packages, our package requires
   846   the \emph{term} of the @{text "quot_thm"} to be given by the user.\footnote{Though we
   833   the \emph{term} of the @{text "quot_thm"} to be given by the user.\footnote{Though we
   847   also provide a fully automated mode, where the @{text "quot_thm"} is guessed
   834   also provide a fully automated mode, where the @{text "quot_thm"} is guessed
   848   from @{text "raw_thm"}.} As a result, it is possible that a user can lift only some
   835   from the form of @{text "raw_thm"}.} As a result, it is possible that a user can lift only some
   849   occurrences of a raw type. 
   836   occurrences of a raw type, but not others. 
   850 
   837 
   851   The second and third proof step will always succeed if the appropriate
   838   The second and third proof step will always succeed if the appropriate
   852   respectfulness and preservation theorems are given. In contrast, the first
   839   respectfulness and preservation theorems are given. In contrast, the first
   853   proof step can fail: a theorem given by the user does not always
   840   proof step can fail: a theorem given by the user does not always
   854   imply a regularized version and a stronger one needs to be proved. This
   841   imply a regularized version and a stronger one needs to be proved. This
   855   is outside of the scope where the quotient package can help. An example
   842   is outside of the scope where the quotient package can help. An example
   856   for the failure is the simple statement for integers @{text "0 \<noteq> 1"}.
   843   for this kind of failure is the simple statement for integers @{text "0 \<noteq> 1"}.
   857   It does not follow by lifting from the fact that @{text "(0, 0) \<noteq> (1, 0)"}. 
   844   One might hope that it can be proved by lifting @{text "(0, 0) \<noteq> (1, 0)"},  
   858   The raw theorem only shows that particular element in the
   845   but the raw theorem only shows that particular element in the
   859   equivalence classes are not equal. A more general statement saying that
   846   equivalence classes are not equal. A more general statement stipulating that
   860   the equivalence classes are not equal is necessary.
   847   the equivalence classes are not equal is necessary, and then leads to the
       
   848   theorem  @{text "0 \<noteq> 1"}.
   861 
   849 
   862   In the following we will first define the statement of the
   850   In the following we will first define the statement of the
   863   regularized theorem based on @{text "raw_thm"} and
   851   regularized theorem based on @{text "raw_thm"} and
   864   @{text "quot_thm"}. Then we define the statement of the injected theorem, based
   852   @{text "quot_thm"}. Then we define the statement of the injected theorem, based
   865   on @{text "reg_thm"} theorem and @{text "quot_thm"}. We then show the 3 proofs,
   853   on @{text "reg_thm"} and @{text "quot_thm"}. We then show the three proof steps,
   866   which can all be performed independently from each other.
   854   which can all be performed independently from each other.
   867 
   855 
   868   We define the function @{text REG}. The intuition
   856   We first define the function @{text REG}. The intuition
   869   behind this function is that it replaces quantifiers and
   857   behind this function is that it replaces quantifiers and
   870   abstractions involving raw types by bounded ones, and equalities
   858   abstractions involving raw types by bounded ones, and equalities
   871   involving raw types are replaced by appropriate aggregate
   859   involving raw types are replaced by appropriate aggregate
   872   equivalence relations. It is defined as follows:
   860   equivalence relations. It is defined as follows:
   873 
   861 
   897   \end{center}
   885   \end{center}
   898   %
   886   %
   899   \noindent
   887   \noindent
   900   In the above definition we omitted the cases for existential quantifiers
   888   In the above definition we omitted the cases for existential quantifiers
   901   and unique existential quantifiers, as they are very similar to the cases
   889   and unique existential quantifiers, as they are very similar to the cases
   902   for the universal quantifier.
   890   for the universal quantifier. For the third and fourt clause, note that 
       
   891   @{text "\<forall>x. P"} is defined as @{text "\<forall> (\<lambda>x. P)"}.
   903 
   892 
   904   Next we define the function @{text INJ} which takes as argument
   893   Next we define the function @{text INJ} which takes as argument
   905   @{text "reg_thm"} and @{text "quot_thm"} (both as
   894   @{text "reg_thm"} and @{text "quot_thm"} (both as
   906   terms) and returns @{text "inj_thm"}:
   895   terms) and returns @{text "inj_thm"}:
   907 
   896 
   939 
   928 
   940   In the first proof step, establishing @{text "raw_thm \<longrightarrow> reg_thm"}, we always 
   929   In the first proof step, establishing @{text "raw_thm \<longrightarrow> reg_thm"}, we always 
   941   start with an implication. Isabelle provides \emph{mono} rules that can split up 
   930   start with an implication. Isabelle provides \emph{mono} rules that can split up 
   942   the implications into simpler implication subgoals. This succeeds for every
   931   the implications into simpler implication subgoals. This succeeds for every
   943   monotone connective, except in places where the function @{text REG} inserted,
   932   monotone connective, except in places where the function @{text REG} inserted,
   944   for example, a quantifier by a bounded quantifier. In this case we have 
   933   for instance, a quantifier by a bounded quantifier. In this case we have 
   945   rules of the form
   934   rules of the form
   946 
   935 
   947   @{text [display, indent=10] "(\<forall>x. R x \<longrightarrow> (P x \<longrightarrow> Q x)) \<longrightarrow> (\<forall>x. P x \<longrightarrow> \<forall>x \<in> R. Q x)"}
   936   @{text [display, indent=10] "(\<forall>x. R x \<longrightarrow> (P x \<longrightarrow> Q x)) \<longrightarrow> (\<forall>x. P x \<longrightarrow> \<forall>x \<in> R. Q x)"}
   948 
   937 
   949   \noindent
   938   \noindent
   959 
   948 
   960   @{thm [display, indent=10] (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
   949   @{thm [display, indent=10] (concl) ball_reg_eqv_range[of R\<^isub>1 R\<^isub>2, no_vars]}
   961 
   950 
   962   \noindent
   951   \noindent
   963   The last theorem is new in comparison with Homeier's package. There the
   952   The last theorem is new in comparison with Homeier's package. There the
   964   injection procedure would be used to prove such goals, and there
   953   injection procedure would be used to prove such goals, and 
   965   the assumption about the equivalence relation would be used. We use the above theorem directly,
   954   the assumption about the equivalence relation would be used. We use the above theorem directly,
   966   because this allows us to completely separate the first and the second
   955   because this allows us to completely separate the first and the second
   967   proof step into independent ``units''.
   956   proof step into two independent ``units''.
   968 
   957 
   969   The second proof step, establishing @{text "reg_thm \<longleftrightarrow> inj_thm"},  starts with an equality. 
   958   The second proof step, establishing @{text "reg_thm \<longleftrightarrow> inj_thm"},  starts with an equality. 
   970   The proof again follows the structure of the
   959   The proof again follows the structure of the
   971   two underlying terms, and is defined for a goal being a relation between these two terms.
   960   two underlying terms, and is defined for a goal being a relation between these two terms.
   972 
   961 
   987   \end{itemize}
   976   \end{itemize}
   988 
   977 
   989   We defined the theorem @{text "inj_thm"} in such a way that
   978   We defined the theorem @{text "inj_thm"} in such a way that
   990   establishing the equivalence @{text "inj_thm \<longleftrightarrow> quot_thm"} can be
   979   establishing the equivalence @{text "inj_thm \<longleftrightarrow> quot_thm"} can be
   991   achieved by rewriting @{text "inj_thm"} with the preservation theorems and quotient
   980   achieved by rewriting @{text "inj_thm"} with the preservation theorems and quotient
   992   definitions. Then all lifted constants, their definitions
   981   definitions. Then for all lifted constants, their definitions
   993   are used to fold the @{term Rep} with the raw constant. Next for
   982   are used to fold the @{term Rep} with the raw constant. Next for
   994   all abstractions and quantifiers the lambda and
   983   all abstractions and quantifiers the lambda and
   995   quantifier preservation theorems are used to replace the
   984   quantifier preservation theorems are used to replace the
   996   variables that include raw types with respects by quantifiers
   985   variables that include raw types with respects by quantifiers
   997   over variables that include quotient types. We show here only
   986   over variables that include quotient types. We show here only
  1051   The user can then specify the constants on the quotient type:
  1040   The user can then specify the constants on the quotient type:
  1052 
  1041 
  1053   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1042   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1054   \begin{tabular}{@ {}l}
  1043   \begin{tabular}{@ {}l}
  1055   \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\[3mm]
  1044   \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0 :: nat, 0 :: nat)"}\\[3mm]
  1056   \isacommand{fun}~~@{text "plus_raw"}~~\isacommand{where}~~%
  1045   \isacommand{fun}~~@{text "add_pair"}~~\isacommand{where}~~%
  1057   @{text "plus_raw (m, n) (p, q) = (m + p :: nat, n + q :: nat)"}\\
  1046   @{text "add_pair (m, n) (p, q) \<equiv> (m + p :: nat, n + q :: nat)"}\\
  1058   \isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
  1047   \isacommand{quotient\_definition}~~@{text "+ :: int \<Rightarrow> int \<Rightarrow> int"}~~%
  1059   \isacommand{is}~~@{text "plus_raw"}\\
  1048   \isacommand{is}~~@{text "add_pair"}\\
  1060   \end{tabular}
  1049   \end{tabular}
  1061   \end{isabelle}
  1050   \end{isabelle}
  1062 
  1051 
  1063   \noindent
  1052   \noindent
  1064   The following theorem about addition on the raw level can be proved.
  1053   The following theorem about addition on the raw level can be proved.
  1065 
  1054 
  1066   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1055   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1067   \isacommand{lemma}~~@{text "plus_zero_raw: int_rel (plus_raw (0, 0) x) x"}
  1056   \isacommand{lemma}~~@{text "add_pair_zero: int_rel (add_pair (0, 0) x) x"}
  1068   \end{isabelle}
  1057   \end{isabelle}
  1069 
  1058 
  1070   \noindent
  1059   \noindent
  1071   If the user attempts to lift this theorem, all proof obligations are 
  1060   If the user attempts to lift this theorem, all proof obligations are 
  1072   automatically discharged, except the respectfulness
  1061   automatically discharged, except the respectfulness
  1073   proof for @{text "plus_raw"}:
  1062   proof for @{text "add_pair"}:
  1074 
  1063 
  1075   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1064   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1076   \begin{tabular}{@ {}l}
  1065   \begin{tabular}{@ {}l}
  1077   \isacommand{lemma}~~@{text "[quot_respect]:\\ 
  1066   \isacommand{lemma}~~@{text "[quot_respect]:"}\\ 
  1078   (int_rel \<doublearr> int_rel \<doublearr> int_rel) plus_raw plus_raw"}
  1067   @{text "(int_rel \<doublearr> int_rel \<doublearr> int_rel) add_pair add_pair"}
  1079   \end{tabular}
  1068   \end{tabular}
  1080   \end{isabelle}
  1069   \end{isabelle}
  1081 
  1070 
  1082   \noindent
  1071   \noindent
  1083   This can be discharged automatically by Isabelle when telling it to unfold the definition
  1072   This can be discharged automatically by Isabelle when telling it to unfold the definition
  1084   of @{text "\<doublearr>"}.
  1073   of @{text "\<doublearr>"}.
  1085   After this, the user can prove the lifted lemma explicitly:
  1074   After this, the user can prove the lifted lemma explicitly:
  1086 
  1075 
  1087   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1076   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1088   \isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting plus_zero_raw"}
  1077   \isacommand{lemma}~~@{text "0 + (x :: int) = x"}~~\isacommand{by}~~@{text "lifting add_pair_zero"}
  1089   \end{isabelle}
  1078   \end{isabelle}
  1090 
  1079 
  1091   \noindent
  1080   \noindent
  1092   or by the completely automated mode by stating:
  1081   or by the completely automated mode by stating:
  1093 
  1082 
  1094   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1083   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %
  1095   \isacommand{thm}~~@{text "plus_zero_raw[quot_lifted]"}
  1084   \isacommand{thm}~~@{text "add_pair_zero[quot_lifted]"}
  1096   \end{isabelle}
  1085   \end{isabelle}
  1097 
  1086 
  1098   \noindent
  1087   \noindent
  1099   Both methods give the same result, namely
  1088   Both methods give the same result, namely
  1100 
  1089 
  1101   @{text [display, indent=10] "0 + x = x"}
  1090   @{text [display, indent=10] "0 + x = x"}
  1102 
  1091 
  1103   \noindent
  1092   \noindent
  1104   Although seemingly simple, arriving at this result without the help of a quotient
  1093   Although seemingly simple, arriving at this result without the help of a quotient
  1105   package requires substantial reasoning effort.
  1094   package requires a substantial reasoning effort.
  1106 *}
  1095 *}
  1107 
  1096 
  1108 section {* Conclusion and Related Work\label{sec:conc}*}
  1097 section {* Conclusion and Related Work\label{sec:conc}*}
  1109 
  1098 
  1110 text {*
  1099 text {*
  1112   The code of the quotient package and the examples described here are
  1101   The code of the quotient package and the examples described here are
  1113   already included in the
  1102   already included in the
  1114   standard distribution of Isabelle.\footnote{Available from
  1103   standard distribution of Isabelle.\footnote{Available from
  1115   \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} It is
  1104   \href{http://isabelle.in.tum.de/}{http://isabelle.in.tum.de/}.} It is
  1116   heavily used in the new version of Nominal Isabelle, which provides a convenient reasoning
  1105   heavily used in the new version of Nominal Isabelle, which provides a convenient reasoning
  1117   infrastructure for programming language calculi involving binders.  
  1106   infrastructure for programming language calculi involving general binders.  
  1118   To achieve this, it builds types representing @{text \<alpha>}-equivalent terms.
  1107   To achieve this, it builds types representing @{text \<alpha>}-equivalent terms.
  1119   Earlier
  1108   Earlier
  1120   versions of Nominal Isabelle have been used successfully in formalisations
  1109   versions of Nominal Isabelle have been used successfully in formalisations
  1121   of an equivalence checking algorithm for LF \cite{UrbanCheneyBerghofer08},
  1110   of an equivalence checking algorithm for LF \cite{UrbanCheneyBerghofer08},
  1122   Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for
  1111   Typed Scheme~\cite{TobinHochstadtFelleisen08}, several calculi for
  1138   Paulson showed a construction of quotients that does not require the
  1127   Paulson showed a construction of quotients that does not require the
  1139   Hilbert Choice operator, but also only first-order theorems can be lifted~\cite{Paulson06}.
  1128   Hilbert Choice operator, but also only first-order theorems can be lifted~\cite{Paulson06}.
  1140   The most related work to our package is the package for HOL4 by Homeier~\cite{Homeier05}.
  1129   The most related work to our package is the package for HOL4 by Homeier~\cite{Homeier05}.
  1141   He introduced most of the abstract notions about quotients and also deals with the
  1130   He introduced most of the abstract notions about quotients and also deals with the
  1142   lifting of higher-order theorems. However, he cannot deal with quotient compositions (needed
  1131   lifting of higher-order theorems. However, he cannot deal with quotient compositions (needed
  1143   for lifting theorems about @{text flat}. Also, a number of his definitions, like @{text ABS},
  1132   for lifting theorems about @{text flat}). Also, a number of his definitions, like @{text ABS},
  1144   @{text REP} and @{text INJ} etc only exist in ML-code. On the other hand, Homeier
  1133   @{text REP} and @{text INJ} etc only exist in \cite{Homeier05} as ML-code, not included
  1145   is able to deal with partial quotient constructions, which we have not implemented.
  1134   in the paper. 
  1146 
  1135 
  1147   One advantage of our package is that it is modular---in the sense that every step
  1136   One advantage of our package is that it is modular---in the sense that every step
  1148   in the quotient construction can be done independently (see the criticism of Paulson
  1137   in the quotient construction can be done independently (see the criticism of Paulson
  1149   about other quotient packages). This modularity is essential in the context of
  1138   about other quotient packages). This modularity is essential in the context of
  1150   Isabelle, which supports type-classes and locales.
  1139   Isabelle, which supports type-classes and locales.
  1151 
  1140 
  1152   Another feature of our quotient package is that when lifting theorems, we can
  1141   Another feature of our quotient package is that when lifting theorems, teh user can
  1153   precisely specify what the lifted theorem should look like. This feature is
  1142   precisely specify what the lifted theorem should look like. This feature is
  1154   necessary, for example, when lifting an inductive principle about two lists.
  1143   necessary, for example, when lifting an induction principle for two lists.
  1155   This principle has as the conclusion a predicate of the form @{text "P xs ys"},
  1144   This principle has as the conclusion a predicate of the form @{text "P xs ys"},
  1156   and we can precisely specify whether we want to quotient @{text "xs"} or @{text "ys"},
  1145   and we can precisely specify whether we want to quotient @{text "xs"} or @{text "ys"},
  1157   or both. We found this feature very useful in the new version of Nominal 
  1146   or both. We found this feature very useful in the new version of Nominal 
  1158   Isabelle.
  1147   Isabelle, where such a choice is required to generate a resoning infrastructure
       
  1148   for alpha-equated terms. 
       
  1149 %%
       
  1150 %% give an example for this
       
  1151 %%
  1159   \medskip
  1152   \medskip
  1160 
  1153 
  1161   \noindent
  1154   \noindent
  1162   {\bf Acknowledgements:} We would like to thank Peter Homeier for the
  1155   {\bf Acknowledgements:} We would like to thank Peter Homeier for the many
  1163   discussions about his HOL4 quotient package and explaining to us
  1156   discussions about his HOL4 quotient package and explaining to us
  1164   some of its finer points in the implementation. Without his patient
  1157   some of its finer points in the implementation. Without his patient
  1165   help, this work would have been impossible.
  1158   help, this work would have been impossible.
  1166 
  1159 
  1167 *}
  1160 *}