# HG changeset patch # User wenzelm # Date 974313822 -3600 # Node ID 4f15b844fea637d8bae49399e8b5b265a31ff284 # Parent 6569febd98e58712b52181209a3de31ef5ffbca8 separate rules for function/operation definitions; diff -r 6569febd98e5 -r 4f15b844fea6 src/HOL/Library/Quotient.thy --- a/src/HOL/Library/Quotient.thy Wed Nov 15 19:42:58 2000 +0100 +++ b/src/HOL/Library/Quotient.thy Wed Nov 15 19:43:42 2000 +0100 @@ -4,7 +4,7 @@ *) header {* - \title{Quotients} + \title{Quotient types} \author{Gertrud Bauer and Markus Wenzel} *} @@ -136,71 +136,114 @@ on quotient types. *} -theorem quot_cond_definition1: +theorem quot_cond_function1: + "(!!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 (rule cong_g) + show "pick \a\ \ a" .. + hence "P (pick \a\) = P a" by (rule cong_P) + also note P + finally show "P (pick \a\)" . + qed + finally show ?thesis . +qed + +theorem quot_function1: + "(!!X. f X == g (pick X)) ==> + (!!x x'. x \ x' ==> g x = g x') ==> + f \a\ = g a" +proof - + case antecedent from this refl TrueI + show ?thesis by (rule quot_cond_function1) +qed + +theorem quot_cond_operation1: "(!!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 . + assume defn: "!!X. f X == \g (pick X)\" + assume "!!x x'. x \ x' ==> P x ==> P x' ==> g x \ g x'" + hence cong_g: "!!x x'. x \ x' ==> P x ==> P x' ==> \g x\ = \g x'\" .. + assume "!!x x'. x \ x' ==> P x = P x'" and "P a" + with defn cong_g show ?thesis by (rule quot_cond_function1) qed -theorem quot_definition1: +theorem quot_operation1: "(!!X. f X == \g (pick X)\) ==> (!!x x'. x \ x' ==> g x \ g x') ==> f \a\ = \g a\" proof - case antecedent from this refl TrueI - show ?thesis by (rule quot_cond_definition1) + show ?thesis by (rule quot_cond_operation1) 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') ==> +theorem quot_cond_function2: + "(!!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\" + 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_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\) (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 + 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 (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 finally show ?thesis . qed -theorem quot_definition2: +theorem quot_function2: + "(!!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" +proof - + case antecedent from this refl TrueI + show ?thesis by (rule quot_cond_function2) +qed + +theorem quot_cond_operation2: + "(!!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 defn: "!!X Y. f X Y == \g (pick X) (pick Y)\" + assume "!!x x' y y'. x \ x' ==> y \ y' ==> P x y ==> P x' y' + ==> g x y \ g x' y'" + hence cong_g: "!!x x' y y'. x \ x' ==> y \ y' ==> P x y ==> P x' y' + ==> \g x y\ = \g x' y'\" .. + assume "!!x x' y y'. x \ x' ==> y \ y' ==> P x y = P x' y'" and "P a b" + with defn cong_g show ?thesis by (rule quot_cond_function2) +qed + +theorem quot_operation2: "(!!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\" proof - case antecedent from this refl TrueI - show ?thesis by (rule quot_cond_definition2) + show ?thesis by (rule quot_cond_operation2) qed text {*