Quotient-Paper/Paper.thy
changeset 2269 e4699a240d2c
parent 2265 9c44db3eef95
child 2270 4ae32dd02d14
equal deleted inserted replaced
2265:9c44db3eef95 2269:e4699a240d2c
     7         "../Nominal/FSet"
     7         "../Nominal/FSet"
     8 begin
     8 begin
     9 
     9 
    10 notation (latex output)
    10 notation (latex output)
    11   rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
    11   rel_conj ("_ \<circ>\<circ>\<circ> _" [53, 53] 52) and
    12   pred_comp ("_ \<circ>\<circ> _") and
    12   pred_comp ("_ \<circ>\<circ> _" [1, 1] 30) and
    13   "op -->" (infix "\<longrightarrow>" 100) and
    13   "op -->" (infix "\<longrightarrow>" 100) and
    14   "==>" (infix "\<Longrightarrow>" 100) and
    14   "==>" (infix "\<Longrightarrow>" 100) and
    15   fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
    15   fun_map ("_ \<^raw:\mbox{\singlearr}> _" 51) and
    16   fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
    16   fun_rel ("_ \<^raw:\mbox{\doublearr}> _" 51) and
    17   list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
    17   list_eq (infix "\<approx>" 50) and (* Not sure if we want this notation...? *)
    88   which states that two lists are equivalent if every element in one list is
    88   which states that two lists are equivalent if every element in one list is
    89   also member in the other. The empty finite set, written @{term "{||}"}, can
    89   also member in the other. The empty finite set, written @{term "{||}"}, can
    90   then be defined as the empty list and the union of two finite sets, written
    90   then be defined as the empty list and the union of two finite sets, written
    91   @{text "\<union>"}, as list append.
    91   @{text "\<union>"}, as list append.
    92 
    92 
    93   Quotients are important in a variety of areas, but they are ubiquitous in
    93   Quotients are important in a variety of areas, but they are really ubiquitous in
    94   the area of reasoning about programming language calculi. A simple example
    94   the area of reasoning about programming language calculi. A simple example
    95   is the lambda-calculus, whose raw terms are defined as
    95   is the lambda-calculus, whose raw terms are defined as
    96 
    96 
    97 
    97 
    98   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
    98   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   152   \end{tikzpicture}
   152   \end{tikzpicture}
   153   \end{center}
   153   \end{center}
   154 
   154 
   155   \noindent
   155   \noindent
   156   The starting point is an existing type, to which we refer as the
   156   The starting point is an existing type, to which we refer as the
   157   \emph{raw type}, over which an equivalence relation given by the user is
   157   \emph{raw type} and over which an equivalence relation given by the user is
   158   defined. With this input the package introduces a new type, to which we
   158   defined. With this input the package introduces a new type, to which we
   159   refer as the \emph{quotient type}. This type comes with an
   159   refer as the \emph{quotient type}. This type comes with an
   160   \emph{abstraction} and a \emph{representation} function, written @{text Abs}
   160   \emph{abstraction} and a \emph{representation} function, written @{text Abs}
   161   and @{text Rep}.\footnote{Actually slightly more basic functions are given;
   161   and @{text Rep}.\footnote{Actually slightly more basic functions are given;
   162   the functions @{text Abs} and @{text Rep} need to be derived from them. We
   162   the functions @{text Abs} and @{text Rep} need to be derived from them. We
   163   will show the details later. } These functions relate elements in the
   163   will show the details later. } They relate elements in the
   164   existing type to elements in the new type and vice versa; they can be uniquely
   164   existing type to elements in the new type and vice versa, and can be uniquely
   165   identified by their quotient type. For example for the integer quotient construction
   165   identified by their quotient type. For example for the integer quotient construction
   166   the types of @{text Abs} and @{text Rep} are
   166   the types of @{text Abs} and @{text Rep} are
   167 
   167 
   168 
   168 
   169   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   169   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   263 *}
   263 *}
   264 
   264 
   265 section {* Preliminaries and General Quotients\label{sec:prelims} *}
   265 section {* Preliminaries and General Quotients\label{sec:prelims} *}
   266 
   266 
   267 text {*
   267 text {*
   268   We describe here briefly the most basic notions of HOL we rely on, and 
   268   We describe in this section briefly the most basic notions of HOL we rely on, and 
   269   the important definitions given by Homeier for quotients \cite{Homeier05}.
   269   the important definitions given by Homeier for quotients \cite{Homeier05}.
   270 
   270 
   271   At its core HOL is based on a simply-typed term language, where types are 
   271   At its core HOL is based on a simply-typed term language, where types are 
   272   recorded in Church-style fashion (that means, we can infer the type of 
   272   recorded in Church-style fashion (that means, we can always infer the type of 
   273   a term and its subterms without any additional information). The grammars
   273   a term and its subterms without any additional information). The grammars
   274   for types and terms are as follows
   274   for types and terms are as follows
   275 
   275 
   276   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   276   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   277   \begin{tabular}{@ {}rl@ {\hspace{3mm}}l@ {}}
   277   \begin{tabular}{@ {}rl@ {\hspace{3mm}}l@ {}}
   283 
   283 
   284   \noindent
   284   \noindent
   285   We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
   285   We often write just @{text \<kappa>} for @{text "() \<kappa>"}, and use @{text "\<alpha>s"} and
   286   @{text "\<sigma>s"} to stand for collections of type variables and types,
   286   @{text "\<sigma>s"} to stand for collections of type variables and types,
   287   respectively.  The type of a term is often made explicit by writing @{text
   287   respectively.  The type of a term is often made explicit by writing @{text
   288   "t :: \<sigma>"}. HOL contains a type @{typ bool} for booleans and the function 
   288   "t :: \<sigma>"}. HOL includes a type @{typ bool} for booleans and the function 
   289   type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains
   289   type, written @{text "\<sigma> \<Rightarrow> \<tau>"}. HOL also contains
   290   many primitive and defined constants; for example equality @{text "= :: \<sigma> \<Rightarrow>
   290   many primitive and defined constants; this includes equality, with type @{text "= :: \<sigma> \<Rightarrow>
   291   \<sigma> \<Rightarrow> bool"} and the identity function @{text "id :: \<sigma> => \<sigma>"} (the former
   291   \<sigma> \<Rightarrow> bool"}, and the identity function, with type @{text "id :: \<sigma> => \<sigma>"} (the former
   292   being primitive and the latter being defined as @{text
   292   being primitive and the latter being defined as @{text
   293   "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}).
   293   "\<lambda>x\<^sup>\<sigma>. x\<^sup>\<sigma>"}).
   294 
   294 
   295   An important point to note is that theorems in HOL can be seen as a subset
   295   An important point to note is that theorems in HOL can be seen as a subset
   296   of terms that are constructed specially (namely through axioms and prove
   296   of terms that are constructed specially (namely through axioms and prove
   297   rules). As a result we are able later on to define automatic proof
   297   rules). As a result we are able to define automatic proof
   298   procedures showing that one theorem implies another by decomposing the term
   298   procedures showing that one theorem implies another by decomposing the term
   299   underlying the first theorem.
   299   underlying the first theorem.
   300 
   300 
   301   Like Homeier, our work relies on map-functions defined for every type constructor,
   301   Like Homeier, our work relies on map-functions defined for every type constructor,
   302   like @{text map} for lists. Homeier describes others for products, sums,
   302   like @{text map} for lists. Homeier describes in \cite{Homeier05} map-functions
       
   303   for products, sums,
   303   options and also the following map for function types
   304   options and also the following map for function types
   304 
   305 
   305   @{thm [display, indent=10] fun_map_def[no_vars, THEN eq_reflection]}
   306   @{thm [display, indent=10] fun_map_def[no_vars, THEN eq_reflection]}
   306 
   307 
   307   \noindent
   308   \noindent
   309   uniform, definition for @{text add} shown in \eqref{adddef}:
   310   uniform, definition for @{text add} shown in \eqref{adddef}:
   310 
   311 
   311   @{text [display, indent=10] "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
   312   @{text [display, indent=10] "add \<equiv> (Rep_int \<singlearr> Rep_int \<singlearr> Abs_int) add_pair"}
   312 
   313 
   313   \noindent
   314   \noindent
   314   We will sometimes refer to a map-function defined for a type-constructor @{text \<kappa>}
   315   Using extensionality and unfolding the definition, we can get back to \eqref{adddef}. 
   315   as @{text "map_\<kappa>"}. 
   316   In what follows we shall use the terminology @{text "map_\<kappa>"} for a map-function 
       
   317   defined for the type-constructor @{text \<kappa>}. In our implementation we have 
       
   318   a database of map-functions that can be dynamically extended.
   316 
   319 
   317   It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
   320   It will also be necessary to have operators, referred to as @{text "rel_\<kappa>"},
   318   which define equivalence relations in terms of constituent equivalence
   321   which define equivalence relations in terms of constituent equivalence
   319   relations. For example given two equivalence relations @{text "R\<^isub>1"}
   322   relations. For example given two equivalence relations @{text "R\<^isub>1"}
   320   and @{text "R\<^isub>2"}, we can define an equivalence relations over 
   323   and @{text "R\<^isub>2"}, we can define an equivalence relations over 
   321   products as follows
   324   products as follows
   322   %
   325   %
   323   @{text [display, indent=10] "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
   326   @{text [display, indent=10] "(R\<^isub>1 \<tripple> R\<^isub>2) (x\<^isub>1, x\<^isub>2) (y\<^isub>1, y\<^isub>2) \<equiv> R\<^isub>1 x\<^isub>1 y\<^isub>1 \<and> R\<^isub>2 x\<^isub>2 y\<^isub>2"}
   324 
   327 
   325   \noindent
   328   \noindent
   326   Similarly, Homeier defines the following operator for defining equivalence 
   329   Homeier gives also the following operator for defining equivalence 
   327   relations over function types:
   330   relations over function types
   328   %
   331   %
   329   @{thm [display, indent=10] fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
   332   @{thm [display, indent=10] fun_rel_def[of "R\<^isub>1" "R\<^isub>2", no_vars, THEN eq_reflection]}
   330 
   333 
   331   The central definition in Homeier's work \cite{Homeier05} relates equivalence 
   334   The central definition in Homeier's work \cite{Homeier05} relates equivalence 
   332   relations, abstraction and representation functions:
   335   relations, abstraction and representation functions:
   341   \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
   344   \item @{thm (rhs3) Quotient_def[of "R", no_vars]}
   342   \end{enumerate}
   345   \end{enumerate}
   343   \end{definition}
   346   \end{definition}
   344 
   347 
   345   \noindent
   348   \noindent
   346   The value of this definition is that validity of @{text Quotient} can be 
   349   The value of this definition is that validity of @{text "Quotient R Abs Rep"} can 
   347   proved in terms of the validity of @{text "Quotient"} over the constituent types. 
   350   often be proved in terms of the validity of @{text "Quotient"} over the constituent 
       
   351   types of @{text "R"}, @{text Abs} and @{text Rep}. 
   348   For example Homeier proves the following property for higher-order quotient
   352   For example Homeier proves the following property for higher-order quotient
   349   types:
   353   types:
   350  
   354  
   351   \begin{proposition}[Function Quotient]\label{funquot}
   355   \begin{proposition}\label{funquot}
   352   @{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2" 
   356   @{thm[mode=IfThen] fun_quotient[where ?R1.0="R\<^isub>1" and ?R2.0="R\<^isub>2" 
   353       and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
   357       and ?abs1.0="Abs\<^isub>1" and ?abs2.0="Abs\<^isub>2" and ?rep1.0="Rep\<^isub>1" and ?rep2.0="Rep\<^isub>2"]}
   354   \end{proposition}
   358   \end{proposition}
   355 
   359 
   356   \noindent
   360   \noindent
   357   We will heavily rely on this part of Homeier's work including an extension 
   361   As a result, Homeier was able to build an automatic prover that can nearly
       
   362   always discharge a proof obligation involving @{text "Quotient"}. Our quotient
       
   363   package makes heavy 
       
   364   use of this part of Homeier's work including an extension 
   358   to deal with compositions of equivalence relations defined as follows
   365   to deal with compositions of equivalence relations defined as follows
   359 
   366 
   360   \begin{definition}[Composition of Relations]
   367   \begin{definition}[Composition of Relations]
   361   @{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
   368   @{abbrev "rel_conj R\<^isub>1 R\<^isub>2"} where @{text "\<circ>\<circ>"} is the predicate
   362   composition defined by the rule
   369   composition defined by the rule
   375 
   382 
   376 text {*
   383 text {*
   377   The first step in a quotient construction is to take a name for the new
   384   The first step in a quotient construction is to take a name for the new
   378   type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
   385   type, say @{text "\<kappa>\<^isub>q"}, and an equivalence relation, say @{text R},
   379   defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
   386   defined over a raw type, say @{text "\<sigma>"}. The type of the equivalence
   380   relation must be of type @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
   387   relation must be @{text "\<sigma> \<Rightarrow> \<sigma> \<Rightarrow> bool"}. The user-visible part of
   381   the declaration is therefore
   388   the quotient type declaration is therefore
   382 
   389 
   383   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   390   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   384   \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}
   391   \isacommand{quotient\_type}~~@{text "\<alpha>s \<kappa>\<^isub>q = \<sigma> / R"}\hfill\numbered{typedecl}
   385   \end{isabelle}
   392   \end{isabelle}
   386 
   393 
   387   \noindent
   394   \noindent
   388   and a proof that @{text "R"} is indeed an equivalence relation. Two concrete
   395   and a proof that @{text "R"} is indeed an equivalence relation. Two concrete
   389   examples are
   396   examples are
   397   \end{isabelle}
   404   \end{isabelle}
   398 
   405 
   399   \noindent
   406   \noindent
   400   which introduce the type of integers and of finite sets using the
   407   which introduce the type of integers and of finite sets using the
   401   equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
   408   equivalence relations @{text "\<approx>\<^bsub>nat \<times> nat\<^esub>"} and @{text
   402   "\<approx>\<^bsub>list\<^esub>"} defined earlier in \eqref{natpairequiv} and
   409   "\<approx>\<^bsub>list\<^esub>"} defined in \eqref{natpairequiv} and
   403   \eqref{listequiv}, respectively (the proofs about being equivalence
   410   \eqref{listequiv}, respectively (the proofs about being equivalence
   404   relations is omitted).  Given this data, we declare internally 
   411   relations is omitted).  Given this data, we declare for \eqref{typedecl} internally 
   405   the quotient types as
   412   the quotient types as
   406   
   413   
   407   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   414   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   408   \isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
   415   \isacommand{typedef}~~@{text "\<alpha>s \<kappa>\<^isub>q = {c. \<exists>x. c = R x}"}
   409   \end{isabelle}
   416   \end{isabelle}
   411   \noindent
   418   \noindent
   412   where the right-hand side is the (non-empty) set of equivalence classes of
   419   where the right-hand side is the (non-empty) set of equivalence classes of
   413   @{text "R"}. The restriction in this declaration is that the type variables
   420   @{text "R"}. The restriction in this declaration is that the type variables
   414   in the raw type @{text "\<sigma>"} must be included in the type variables @{text
   421   in the raw type @{text "\<sigma>"} must be included in the type variables @{text
   415   "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will provide us with the following
   422   "\<alpha>s"} declared for @{text "\<kappa>\<^isub>q"}. HOL will provide us with the following
   416   abstraction and representation functions having the type
   423   abstraction and representation functions 
   417 
   424 
   418   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   425   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   419   @{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"}
   426   @{text "abs_\<kappa>\<^isub>q :: \<sigma> set \<Rightarrow> \<alpha>s \<kappa>\<^isub>q"}\hspace{10mm}@{text "rep_\<kappa>\<^isub>q :: \<alpha>s \<kappa>\<^isub>q \<Rightarrow> \<sigma> set"}
   420   \end{isabelle}
   427   \end{isabelle}
   421 
   428 
   422   \noindent 
   429   \noindent 
   423   They relate the new quotient type and equivalence classes of the raw
   430   As can be seen from the type, they relate the new quotient type and equivalence classes of the raw
   424   type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
   431   type. However, as Homeier \cite{Homeier05} noted, it is much more convenient
   425   to work with the following derived abstraction and representation functions
   432   to work with the following derived abstraction and representation functions
   426 
   433 
   427   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   434   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   428   @{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)"}
   435   @{text "Abs_\<kappa>\<^isub>q x \<equiv> abs_\<kappa>\<^isub>q (R x)"}\hspace{10mm}@{text "Rep_\<kappa>\<^isub>q x \<equiv> \<epsilon> (rep_\<kappa>\<^isub>q x)"}
   445   as above (for the proof see \cite{Homeier05}).
   452   as above (for the proof see \cite{Homeier05}).
   446 
   453 
   447   The next step in a quotient construction is to introduce definitions of new constants
   454   The next step in a quotient construction is to introduce definitions of new constants
   448   involving the quotient type. These definitions need to be given in terms of concepts
   455   involving the quotient type. These definitions need to be given in terms of concepts
   449   of the raw type (remember this is the only way how to extend HOL
   456   of the raw type (remember this is the only way how to extend HOL
   450   with new definitions). For the user visible is the declaration
   457   with new definitions). For the user the visible part of such definitions is the declaration
   451 
   458 
   452   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   459   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   453   \isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
   460   \isacommand{quotient\_definition}~~@{text "c :: \<tau>"}~~\isacommand{is}~~@{text "t :: \<sigma>"}
   454   \end{isabelle}
   461   \end{isabelle}
   455 
   462 
   456   \noindent
   463   \noindent
   457   where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
   464   where @{text t} is the definiens (its type @{text \<sigma>} can always be inferred)
   458   and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
   465   and @{text "c"} is the name of definiendum, whose type @{text "\<tau>"} needs to be
   459   given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ 
   466   given explicitly (the point is that @{text "\<tau>"} and @{text "\<sigma>"} can only differ 
   460   in places where a quotient and raw type are involved). Two concrete examples are
   467   in places where a quotient and raw type is involved). Two concrete examples are
   461 
   468 
   462   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   469   \begin{isabelle}\ \ \ \ \ \ \ \ \ \ %%%
   463   \begin{tabular}{@ {}l}
   470   \begin{tabular}{@ {}l}
   464   \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
   471   \isacommand{quotient\_definition}~~@{text "0 :: int"}~~\isacommand{is}~~@{text "(0::nat, 0::nat)"}\\
   465   \isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
   472   \isacommand{quotient\_definition}~~@{text "\<Union> :: (\<alpha> fset) fset \<Rightarrow> \<alpha> fset"}~~%
   478   abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
   485   abstraction} and \emph{representation functions} using the functions @{text "ABS (\<sigma>,
   479   \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we give below. The idea behind
   486   \<tau>)"} and @{text "REP (\<sigma>, \<tau>)"} whose clauses we give below. The idea behind
   480   these two functions is to recursively descend into the raw types @{text \<sigma>} and 
   487   these two functions is to recursively descend into the raw types @{text \<sigma>} and 
   481   quotient types @{text \<tau>}, and generate the appropriate
   488   quotient types @{text \<tau>}, and generate the appropriate
   482   @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
   489   @{text "Abs"} and @{text "Rep"} in places where the types differ. Therefore
   483   we generate just the identity whenever the types are equal. All clauses
   490   we generate just the identity whenever the types are equal. On the ``way'' down,
   484   are as follows:
   491   however we might have to use map-functions to let @{text Abs} and @{text Rep} act
       
   492   over the appropriate types. The clauses of @{text ABS} and @{text REP} 
       
   493   are as follows (where we use the short-hand notation @{text "ABS (\<sigma>s, \<tau>s)"} to mean 
       
   494   @{text "ABS (\<sigma>\<^isub>1, \<tau>\<^isub>1)\<dots>ABS (\<sigma>\<^isub>i, \<tau>\<^isub>i)"}; similarly for @{text REP}):
   485 
   495 
   486   \begin{center}
   496   \begin{center}
   487   \hfill
   497   \hfill
   488   \begin{tabular}{rcl}
   498   \begin{tabular}{rcl}
   489   \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
   499   \multicolumn{3}{@ {\hspace{-4mm}}l}{equal types:}\\ 
   500   @{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"}
   510   @{text "REP (\<sigma>s \<kappa>, \<tau>s \<kappa>\<^isub>q)"} & $\dn$ & @{text "(MAP(\<rho>s \<kappa>) (REP (\<sigma>s', \<tau>s))) \<circ> Rep_\<kappa>\<^isub>q"}
   501   \end{tabular}\hfill\numbered{ABSREP}
   511   \end{tabular}\hfill\numbered{ABSREP}
   502   \end{center}
   512   \end{center}
   503   %
   513   %
   504   \noindent
   514   \noindent
   505   where in the last two clauses we have that the quotient type @{text "\<alpha>s
   515   where in the last two clauses we have that the type @{text "\<alpha>s
   506   \<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
   516   \<kappa>\<^isub>q"} is the quotient of the raw type @{text "\<rho>s \<kappa>"} (for example
   507   @{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
   517   @{text "int"} and @{text "nat \<times> nat"}, or @{text "\<alpha> fset"} and @{text "\<alpha>
   508   list"}). The quotient construction ensures that the type variables in @{text
   518   list"}). The quotient construction ensures that the type variables in @{text
   509   "\<rho>s"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
   519   "\<rho>s"} must be among the @{text "\<alpha>s"}. The @{text "\<sigma>s'"} are given by the
   510   matchers for the @{text "\<alpha>s"} when matching @{text "\<rho>s \<kappa>"} against
   520   matchers for the @{text "\<alpha>s"} when matching @{text "\<rho>s \<kappa>"} against
   532   aggregate map-function:
   542   aggregate map-function:
   533 
   543 
   534   @{text [display, indent=10] "\<lambda>a b. map_prod (map a) b"}
   544   @{text [display, indent=10] "\<lambda>a b. map_prod (map a) b"}
   535   
   545   
   536   \noindent
   546   \noindent
   537   which we need to define the aggregate abstraction and representation functions.
   547   which we need in order to define the aggregate abstraction and representation 
       
   548   functions.
   538   
   549   
   539   To see how these definitions pan out in practise, let us return to our
   550   To see how these definitions pan out in practise, let us return to our
   540   example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
   551   example about @{term "concat"} and @{term "fconcat"}, where we have the raw type
   541   @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
   552   @{text "(\<alpha> list) list \<Rightarrow> \<alpha> list"} and the quotient type @{text "(\<alpha> fset) fset \<Rightarrow> \<alpha>
   542   fset"}. Feeding them into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
   553   fset"}. Feeding them into @{text ABS} gives us (after some @{text "\<beta>"}-simplifications)
   579   then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
   590   then @{text "Abs"} is of type @{text "\<sigma> \<Rightarrow> \<tau>"} and @{text "Rep"} of type
   580   @{text "\<tau> \<Rightarrow> \<sigma>"}.
   591   @{text "\<tau> \<Rightarrow> \<sigma>"}.
   581   \end{lemma}
   592   \end{lemma}
   582 
   593 
   583   \begin{proof}
   594   \begin{proof}
   584   By induction and analysing the definitions of @{text "ABS"}, @{text "REP"}
   595   By mutual induction and analysing the definitions of @{text "ABS"}, @{text "REP"}
   585   and @{text "MAP"}. The cases of equal types and function types are
   596   and @{text "MAP"}. The cases of equal types and function types are
   586   straightforward (the latter follows from @{text "\<singlearr>"} having the
   597   straightforward (the latter follows from @{text "\<singlearr>"} having the
   587   type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
   598   type @{text "(\<alpha> \<Rightarrow> \<beta>) \<Rightarrow> (\<gamma> \<Rightarrow> \<delta>) \<Rightarrow> (\<beta> \<Rightarrow> \<gamma>) \<Rightarrow> (\<alpha> \<Rightarrow> \<delta>)"}). In case of equal type
   588   constructors we can observe that a map-function after applying the functions
   599   constructors we can observe that a map-function after applying the functions
   589   @{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}.  The
   600   @{text "ABS (\<sigma>s, \<tau>s)"} produces a term of type @{text "\<sigma>s \<kappa> \<Rightarrow> \<tau>s \<kappa>"}.  The