# HG changeset patch # User wenzelm # Date 1419766623 -3600 # Node ID a1d6d6db781be34901b77cc4d486321a0154e7ad # Parent 682aa538c5c088a2c1461cfacf6728023abd3d03 modernized historic example; diff -r 682aa538c5c0 -r a1d6d6db781b src/HOL/Library/Quotient_Type.thy --- a/src/HOL/Library/Quotient_Type.thy Sun Dec 28 12:18:01 2014 +0100 +++ b/src/HOL/Library/Quotient_Type.thy Sun Dec 28 12:37:03 2014 +0100 @@ -2,63 +2,65 @@ Author: Markus Wenzel, TU Muenchen *) -section {* Quotient types *} +section \Quotient types\ theory Quotient_Type imports Main begin -text {* - We introduce the notion of quotient types over equivalence relations - via type classes. -*} +text \We introduce the notion of quotient types over equivalence relations + via type classes.\ + -subsection {* Equivalence relations and quotient types *} +subsection \Equivalence relations and quotient types\ -text {* - \medskip Type class @{text equiv} models equivalence relations @{text - "\ :: 'a => 'a => bool"}. -*} +text \Type class @{text equiv} models equivalence relations + @{text "\ :: 'a \ 'a \ bool"}.\ class eqv = - fixes eqv :: "'a \ 'a \ bool" (infixl "\" 50) + fixes eqv :: "'a \ 'a \ bool" (infixl "\" 50) class equiv = eqv + assumes equiv_refl [intro]: "x \ x" - assumes equiv_trans [trans]: "x \ y \ y \ z \ x \ z" - assumes equiv_sym [sym]: "x \ y \ y \ x" + and equiv_trans [trans]: "x \ y \ y \ z \ x \ z" + and equiv_sym [sym]: "x \ y \ y \ x" +begin -lemma equiv_not_sym [sym]: "\ (x \ y) ==> \ (y \ (x::'a::equiv))" +lemma equiv_not_sym [sym]: "\ x \ y \ \ y \ x" proof - - assume "\ (x \ y)" then show "\ (y \ x)" - by (rule contrapos_nn) (rule equiv_sym) + assume "\ x \ y" + then show "\ y \ x" by (rule contrapos_nn) (rule equiv_sym) qed -lemma not_equiv_trans1 [trans]: "\ (x \ y) ==> y \ z ==> \ (x \ (z::'a::equiv))" +lemma not_equiv_trans1 [trans]: "\ x \ y \ y \ z \ \ x \ z" proof - - assume "\ (x \ y)" and "y \ z" - show "\ (x \ z)" + assume "\ x \ y" and "y \ z" + show "\ x \ z" proof assume "x \ z" - also from `y \ z` have "z \ y" .. + also from \y \ z\ have "z \ y" .. finally have "x \ y" . - with `\ (x \ y)` show False by contradiction + with \\ x \ y\ show False by contradiction qed qed -lemma not_equiv_trans2 [trans]: "x \ y ==> \ (y \ z) ==> \ (x \ (z::'a::equiv))" +lemma not_equiv_trans2 [trans]: "x \ y \ \ y \ z \ \ x \ z" proof - - assume "\ (y \ z)" then have "\ (z \ y)" .. - also assume "x \ y" then have "y \ x" .. - finally have "\ (z \ x)" . then show "(\ x \ z)" .. + assume "\ y \ z" + then have "\ z \ y" .. + also + assume "x \ y" + then have "y \ x" .. + finally have "\ z \ x" . + then show "\ x \ z" .. qed -text {* - \medskip The quotient type @{text "'a quot"} consists of all - \emph{equivalence classes} over elements of the base type @{typ 'a}. -*} +end -definition "quot = {{x. a \ x} | a::'a::eqv. True}" +text \The quotient type @{text "'a quot"} consists of all \emph{equivalence + classes} over elements of the base type @{typ 'a}.\ + +definition (in eqv) "quot = {{x. a \ x} | a. True}" typedef 'a quot = "quot :: 'a::eqv set set" unfolding quot_def by blast @@ -66,38 +68,38 @@ lemma quotI [intro]: "{x. a \ x} \ quot" unfolding quot_def by blast -lemma quotE [elim]: "R \ quot ==> (!!a. R = {x. a \ x} ==> C) ==> C" - unfolding quot_def by blast +lemma quotE [elim]: + assumes "R \ quot" + obtains a where "R = {x. a \ x}" + using assms unfolding quot_def by blast -text {* - \medskip Abstracted equivalence classes are the canonical - representation of elements of a quotient type. -*} +text \Abstracted equivalence classes are the canonical representation of + elements of a quotient type.\ -definition - "class" :: "'a::equiv => 'a quot" ("\_\") where - "\a\ = Abs_quot {x. a \ x}" +definition "class" :: "'a::equiv \ 'a quot" ("\_\") + where "\a\ = Abs_quot {x. a \ x}" theorem quot_exhaust: "\a. A = \a\" proof (cases A) - fix R assume R: "A = Abs_quot R" - assume "R \ quot" then have "\a. R = {x. a \ x}" by blast + fix R + assume R: "A = Abs_quot R" + assume "R \ quot" + then have "\a. R = {x. a \ x}" by blast with R have "\a. A = Abs_quot {x. a \ x}" by blast then show ?thesis unfolding class_def . qed -lemma quot_cases [cases type: quot]: "(!!a. A = \a\ ==> C) ==> C" +lemma quot_cases [cases type: quot]: + obtains a where "A = \a\" using quot_exhaust by blast -subsection {* Equality on quotients *} +subsection \Equality on quotients\ -text {* - Equality of canonical quotient elements coincides with the original - relation. -*} +text \Equality of canonical quotient elements coincides with the original + relation.\ -theorem quot_equality [iff?]: "(\a\ = \b\) = (a \ b)" +theorem quot_equality [iff?]: "\a\ = \b\ \ a \ b" proof assume eq: "\a\ = \b\" show "a \ b" @@ -131,11 +133,10 @@ qed -subsection {* Picking representing elements *} +subsection \Picking representing elements\ -definition - pick :: "'a::equiv quot => 'a" where - "pick A = (SOME a. A = \a\)" +definition pick :: "'a::equiv quot \ 'a" + where "pick A = (SOME a. A = \a\)" theorem pick_equiv [intro]: "pick \a\ \ a" proof (unfold pick_def) @@ -143,7 +144,8 @@ proof (rule someI2) show "\a\ = \a\" .. fix x assume "\a\ = \x\" - then have "a \ x" .. then show "x \ a" .. + then have "a \ x" .. + then show "x \ a" .. qed qed @@ -155,17 +157,14 @@ with a show ?thesis by simp qed -text {* - \medskip The following rules support canonical function definitions - on quotient types (with up to two arguments). Note that the - stripped-down version without additional conditions is sufficient - most of the time. -*} +text \The following rules support canonical function definitions on quotient + types (with up to two arguments). Note that the stripped-down version + without additional conditions is sufficient most of the time.\ theorem quot_cond_function: - assumes eq: "!!X Y. P X Y ==> f X Y == g (pick X) (pick Y)" - and cong: "!!x x' y y'. \x\ = \x'\ ==> \y\ = \y'\ - ==> P \x\ \y\ ==> P \x'\ \y'\ ==> g x y = g x' y'" + assumes eq: "\X Y. P X Y \ f X Y \ g (pick X) (pick Y)" + and cong: "\x x' y y'. \x\ = \x'\ \ \y\ = \y'\ + \ P \x\ \y\ \ P \x'\ \y'\ \ g x y = g x' y'" and P: "P \a\ \b\" shows "f \a\ \b\ = g a b" proof - @@ -183,15 +182,15 @@ qed theorem quot_function: - assumes "!!X Y. f X Y == g (pick X) (pick Y)" - and "!!x x' y y'. \x\ = \x'\ ==> \y\ = \y'\ ==> g x y = g x' y'" + assumes "\X Y. f X Y \ g (pick X) (pick Y)" + and "\x x' y y'. \x\ = \x'\ \ \y\ = \y'\ \ g x y = g x' y'" shows "f \a\ \b\ = g a b" using assms and TrueI by (rule quot_cond_function) theorem quot_function': - "(!!X Y. f X Y == g (pick X) (pick Y)) ==> - (!!x x' y y'. x \ x' ==> y \ y' ==> g x y = g x' y') ==> + "(\X Y. f X Y \ g (pick X) (pick Y)) \ + (\x x' y y'. x \ x' \ y \ y' \ g x y = g x' y') \ f \a\ \b\ = g a b" by (rule quot_function) (simp_all only: quot_equality)