# HG changeset patch # User wenzelm # Date 974037026 -3600 # Node ID df3cd3e760466e9d0bc6eb7488fa74effb256c15 # Parent df4e182c0fcd1190fbe4b2cb19e2a1334ae0ce14 quot_cond_definition; diff -r df4e182c0fcd -r df3cd3e76046 src/HOL/Library/Quotient.thy --- a/src/HOL/Library/Quotient.thy Sun Nov 12 14:49:37 2000 +0100 +++ b/src/HOL/Library/Quotient.thy Sun Nov 12 14:50:26 2000 +0100 @@ -73,7 +73,7 @@ relation. *} -theorem equivalence_class_eq [iff?]: "(\a\ = \b\) = (a \ b)" +theorem equivalence_class_iff [iff?]: "(\a\ = \b\) = (a \ b)" proof assume eq: "\a\ = \b\" show "a \ b" @@ -136,19 +136,59 @@ on quotient types. *} +theorem quot_cond_definition1: + "(!!X. f X == \g (pick X)\) ==> + (!!x x'. x \ x' ==> P x ==> P x' ==> g x \ g x') ==> + (!!x x'. x \ x' ==> P x = P x') ==> + P a ==> f \a\ = \g a\" +proof - + assume cong_g: "!!x x'. x \ x' ==> P x ==> P x' ==> g x \ g x'" + assume cong_P: "!!x x'. x \ x' ==> P x = P x'" + assume P: "P a" + assume "!!X. f X == \g (pick X)\" + hence "f \a\ = \g (pick \a\)\" by (simp only:) + also have "\ = \g a\" + proof + show "g (pick \a\) \ g a" + proof (rule cong_g) + show "pick \a\ \ a" .. + hence "P (pick \a\) = P a" by (rule cong_P) + also show "P a" . + finally show "P (pick \a\)" . + qed + qed + finally show ?thesis . +qed + theorem quot_definition1: "(!!X. f X == \g (pick X)\) ==> (!!x x'. x \ x' ==> g x \ g x') ==> f \a\ = \g a\" proof - - assume cong: "!!x x'. x \ x' ==> g x \ g x'" - assume "!!X. f X == \g (pick X)\" - hence "f \a\ = \g (pick \a\)\" by (simp only:) - also have "\ = \g a\" + case antecedent from this refl TrueI + show ?thesis by (rule quot_cond_definition1) +qed + +theorem quot_cond_definition2: + "(!!X Y. f X Y == \g (pick X) (pick Y)\) ==> + (!!x x' y y'. x \ x' ==> y \ y' ==> P x y ==> P x' y' ==> g x y \ g x' y') ==> + (!!x x' y y'. x \ x' ==> y \ y' ==> P x y = P x' y') ==> + P a b ==> f \a\ \b\ = \g a b\" +proof - + assume cong_g: "!!x x' y y'. x \ x' ==> y \ y' ==> P x y ==> P x' y' ==> g x y \ g x' y'" + assume cong_P: "!!x x' y y'. x \ x' ==> y \ y' ==> P x y = P x' y'" + assume P: "P a b" + assume "!!X Y. f X Y == \g (pick X) (pick Y)\" + hence "f \a\ \b\ = \g (pick \a\) (pick \b\)\" by (simp only:) + also have "\ = \g a b\" proof - show "g (pick \a\) \ g a" - proof (rule cong) + show "g (pick \a\) (pick \b\) \ g a b" + proof (rule cong_g) show "pick \a\ \ a" .. + moreover show "pick \b\ \ b" .. + ultimately have "P (pick \a\) (pick \b\) = P a b" by (rule cong_P) + also show "P a b" . + finally show "P (pick \a\) (pick \b\)" . qed qed finally show ?thesis . @@ -159,21 +199,10 @@ (!!x x' y y'. x \ x' ==> y \ y' ==> g x y \ g x' y') ==> f \a\ \b\ = \g a b\" proof - - assume cong: "!!x x' y y'. x \ x' ==> y \ y' ==> g x y \ g x' y'" - assume "!!X Y. f X Y == \g (pick X) (pick Y)\" - hence "f \a\ \b\ = \g (pick \a\) (pick \b\)\" by (simp only:) - also have "\ = \g a b\" - proof - show "g (pick \a\) (pick \b\) \ g a b" - proof (rule cong) - show "pick \a\ \ a" .. - show "pick \b\ \ b" .. - qed - qed - finally show ?thesis . + case antecedent from this refl TrueI + show ?thesis by (rule quot_cond_definition2) qed - text {* \medskip HOL's collection of overloaded standard operations is lifted to quotient types in the canonical manner. @@ -186,6 +215,7 @@ instance quot :: (inverse) inverse .. instance quot :: (power) power .. instance quot :: (number) number .. +instance quot :: (ord) ord .. defs (overloaded) zero_quot_def: "0 == \0\" @@ -198,5 +228,7 @@ divide_quot_def: "X / Y == \pick X / pick Y\" power_quot_def: "X^n == \(pick X)^n\" number_of_quot_def: "number_of b == \number_of b\" + le_quot_def: "X \ Y == pick X \ pick Y" + less_quot_def: "X < Y == pick X < pick Y" end