diff -r f49b5dfabd59 -r 22c6b6144abd Quotient-Paper/Paper.thy --- a/Quotient-Paper/Paper.thy Sun Jun 13 13:42:37 2010 +0200 +++ b/Quotient-Paper/Paper.thy Sun Jun 13 14:39:55 2010 +0200 @@ -332,8 +332,8 @@ @{text "ABS (\, \)"} & $\dn$ & @{text "id"}\\ @{text "REP (\, \)"} & $\dn$ & @{text "id"}\smallskip\\ \multicolumn{3}{@ {\hspace{-4mm}}l}{function types:}\\ - @{text "ABS (\\<^isub>1 \ \\<^isub>2, \\<^isub>1 \ \\<^isub>2)"} & $\dn$ & @{text "REP (\\<^isub>1, \\<^isub>1) \ ABS (\\<^isub>2, \\<^isub>2)"}\\ - @{text "REP (\\<^isub>1 \ \\<^isub>2, \\<^isub>1 \ \\<^isub>2)"} & $\dn$ & @{text "ABS (\\<^isub>1, \\<^isub>1) \ REP (\\<^isub>2, \\<^isub>2)"}\smallskip\\ + @{text "ABS (\\<^isub>1 \ \\<^isub>2, \\<^isub>1 \ \\<^isub>2)"} & $\dn$ & @{text "REP (\\<^isub>1, \\<^isub>1) \ ABS (\\<^isub>2, \\<^isub>2)"}\\ + @{text "REP (\\<^isub>1 \ \\<^isub>2, \\<^isub>1 \ \\<^isub>2)"} & $\dn$ & @{text "ABS (\\<^isub>1, \\<^isub>1) \ REP (\\<^isub>2, \\<^isub>2)"}\smallskip\\ \multicolumn{3}{@ {\hspace{-4mm}}l}{equal type constructors:}\\ @{text "ABS (\s \, \s \)"} & $\dn$ & @{text "map_\ (ABS (\s, \s))"}\\ @{text "REP (\s \, \s \)"} & $\dn$ & @{text "map_\ (REP (\s, \s))"}\smallskip\\ @@ -356,17 +356,42 @@ \begin{center} \begin{tabular}{rcl} - @{text "MAP' (\)"} & $\dn$ & @{text "\"}\\ + @{text "MAP' (\)"} & $\dn$ & @{text "a"}\\ + @{text "MAP' (\)"} & $\dn$ & @{text "id"}\\ @{text "MAP' (\s \)"} & $\dn$ & @{text "map_\ (MAP'(\s))"}\smallskip\\ - @{text "MAP (\)"} & $\dn$ & @{text "\\s. MAP'(\)"} + @{text "MAP (\)"} & $\dn$ & @{text "\as. MAP'(\)"} \end{tabular} \end{center} \noindent - In this definition we abuse the fact that we can interpret type-variables as - term variables, and in the last clause build the abstraction over all + In this definition we abuse the fact that we can interpret type-variables @{text \} as + term variables @{text a}, and in the last clause build an abstraction over all term-variables inside the aggregate map-function generated by @{text "MAP'"}. + This aggregate map-function is necessary if we build quotients, say @{text "(\, \) foo"}, + out of compound raw types, say @{text "(\ list) \ \"}. In this case @{text MAP} + generates the aggregate map-function: + @{text [display, indent=10] "\a b. map_prod (map a) b"} + + \noindent + Returning to our example about @{term "concat"} and @{term "fconcat"} which have the + types @{text "(\ list) list \ \ list"} and @{text "(\ fset) fset \ \ fset"}. Feeding this + into @{text ABS} gives us the abstraction function: + + @{text [display, indent=10] "(map (map id \ Rep_fset) \ Rep_fset) \ Abs_fset \ map id"} + + \noindent + where we already performed some @{text "\"}-simplifications. In our implementation + we further simplified this by noticing the usual laws about @{text "map"}s and @{text "id"}, + namely @{term "map id = id"} and + @{text "f \ id = id \ f = f"}. This gives us the simplified abstarction function + + @{text [display, indent=10] "(map Rep_fset \ Rep_fset) \ Abs_fset"} + + \noindent + which we can use for defining @{term "fconcat"} as follows + + @{text [display, indent=10] "\ \ ((map Rep_fset \ Rep_fset) \ Abs_fset) flat"} Apart from the last 2 points the definition is same as the one implemented in in Homeier's HOL package. Adding composition in last two cases is necessary @@ -394,6 +419,9 @@ then @{text "Abs"} is of type @{text "\ \ \"} and @{text "Rep"} of type @{text "\ \ \"}. \end{lemma} + + This lemma fails in the over-simplified abstraction and representation function used + in Homeier's quotient package. *} subsection {* Respectfulness *}